more modernisation of syntax
authorpaulson <lp15@cam.ac.uk>
Tue, 27 Sep 2022 17:03:23 +0100
changeset 76214 0c18df79b1c8
parent 76213 e44d86131648
child 76215 a642599ffdea
more modernisation of syntax
src/ZF/AC/AC15_WO6.thy
src/ZF/AC/AC16_WO4.thy
src/ZF/AC/AC16_lemmas.thy
src/ZF/AC/AC17_AC1.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/WO1_AC.thy
src/ZF/AC/WO1_WO7.thy
src/ZF/AC/WO2_AC16.thy
src/ZF/AC/WO6_WO1.thy
src/ZF/Arith.thy
src/ZF/ArithSimp.thy
src/ZF/Bin.thy
src/ZF/Cardinal.thy
src/ZF/CardinalArith.thy
src/ZF/Coind/ECR.thy
src/ZF/Coind/Static.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/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/Separation.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/Denotation.thy
src/ZF/Induct/Binary_Trees.thy
src/ZF/Induct/Comb.thy
src/ZF/Induct/FoldSet.thy
src/ZF/Induct/ListN.thy
src/ZF/Induct/Multiset.thy
src/ZF/Induct/Mutil.thy
src/ZF/Induct/PropLog.thy
src/ZF/Induct/Term.thy
src/ZF/InfDatatype.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/Reduction.thy
src/ZF/Resid/Residuals.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/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/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/AC15_WO6.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/AC15_WO6.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -16,7 +16,7 @@
 
 So we don't have to prove all implications of both cases.
 Moreover we don't need to prove AC13(1) \<Longrightarrow> AC1 and AC11 \<Longrightarrow> AC14 as
-Rubin & Rubin do.
+Rubin \<and> Rubin do.
 *)
 
 theory AC15_WO6
@@ -46,7 +46,7 @@
  apply (blast intro: lepoll_Sigma [THEN lepoll_Finite])+
 done
 
-lemma lemma1: "\<lbrakk>\<Union>(C)=A; a \<in> A\<rbrakk> \<Longrightarrow> \<exists>B \<in> C. a \<in> B & B \<subseteq> A"
+lemma lemma1: "\<lbrakk>\<Union>(C)=A; a \<in> A\<rbrakk> \<Longrightarrow> \<exists>B \<in> C. a \<in> B \<and> B \<subseteq> A"
 by fast
 
 lemma lemma2: 
@@ -54,10 +54,10 @@
 by (unfold pairwise_disjoint_def, blast) 
 
 lemma lemma3: 
-     "\<forall>B \<in> {cons(0, x*nat). x \<in> A}. pairwise_disjoint(f`B) &   
-             sets_of_size_between(f`B, 2, n) & \<Union>(f`B)=B   
-     \<Longrightarrow> \<forall>B \<in> A. \<exists>! u. u \<in> f`cons(0, B*nat) & u \<subseteq> cons(0, B*nat) &   
-             0 \<in> u & 2 \<lesssim> u & u \<lesssim> n"
+     "\<forall>B \<in> {cons(0, x*nat). x \<in> A}. pairwise_disjoint(f`B) \<and>   
+             sets_of_size_between(f`B, 2, n) \<and> \<Union>(f`B)=B   
+     \<Longrightarrow> \<forall>B \<in> A. \<exists>! u. u \<in> f`cons(0, B*nat) \<and> u \<subseteq> cons(0, B*nat) \<and>   
+             0 \<in> u \<and> 2 \<lesssim> u \<and> u \<lesssim> n"
 apply (unfold sets_of_size_between_def)
 apply (rule ballI)
 apply (erule_tac x="cons(0, B*nat)" in ballE)
@@ -67,7 +67,7 @@
 lemma lemma4: "\<lbrakk>A \<lesssim> i; Ord(i)\<rbrakk> \<Longrightarrow> {P(a). a \<in> A} \<lesssim> i"
 apply (unfold lepoll_def)
 apply (erule exE)
-apply (rule_tac x = "\<lambda>x \<in> RepFun(A,P). \<mu> j. \<exists>a\<in>A. x=P(a) & f`a=j" 
+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 (erule RepFunE)
@@ -103,10 +103,10 @@
 
 lemma ex_fun_AC13_AC15:
      "\<lbrakk>\<forall>B \<in> {cons(0, x*nat). x \<in> A}.   
-                pairwise_disjoint(f`B) &   
-                sets_of_size_between(f`B, 2, succ(n)) & \<Union>(f`B)=B; 
+                pairwise_disjoint(f`B) \<and>   
+                sets_of_size_between(f`B, 2, succ(n)) \<and> \<Union>(f`B)=B; 
          n \<in> nat\<rbrakk>   
-      \<Longrightarrow> \<exists>f. \<forall>B \<in> A. f`B \<noteq> 0 & f`B \<subseteq> B & f`B \<lesssim> n"
+      \<Longrightarrow> \<exists>f. \<forall>B \<in> A. f`B \<noteq> 0 \<and> f`B \<subseteq> B \<and> f`B \<lesssim> n"
 by (fast del: subsetI notI
          dest!: lemma3 theI intro!: lemma5_1 lemma5_2 lemma5_3)
 
@@ -147,7 +147,7 @@
 by (fast intro!: ltI dest!: ltD)
 
 lemma AC15_WO6_aux1:
-     "\<forall>x \<in> Pow(A)-{0}. f`x\<noteq>0 & f`x \<subseteq> x & f`x \<lesssim> m 
+     "\<forall>x \<in> Pow(A)-{0}. f`x\<noteq>0 \<and> f`x \<subseteq> x \<and> f`x \<lesssim> m 
       \<Longrightarrow> (\<Union>i<\<mu> x. HH(f,A,x)={A}. HH(f,A,i)) = A"
 apply (simp add: Ord_Least [THEN OUN_eq_UN])
 apply (rule equalityI)
@@ -157,7 +157,7 @@
 done
 
 lemma AC15_WO6_aux2:
-     "\<forall>x \<in> Pow(A)-{0}. f`x\<noteq>0 & f`x \<subseteq> x & f`x \<lesssim> m 
+     "\<forall>x \<in> Pow(A)-{0}. f`x\<noteq>0 \<and> f`x \<subseteq> x \<and> f`x \<lesssim> m 
       \<Longrightarrow> \<forall>x < (\<mu> x. HH(f,A,x)={A}). HH(f,A,x) \<lesssim> m"
 apply (rule oallI)
 apply (drule ltD [THEN less_Least_subset_x])
@@ -251,7 +251,7 @@
 by (unfold AC13_def AC14_def AC15_def, fast)
 
 (* ********************************************************************** *)
-(* The redundant proofs; however cited by Rubin & Rubin                   *)
+(* The redundant proofs; however cited by Rubin \<and> Rubin                   *)
 (* ********************************************************************** *)
 
 (* ********************************************************************** *)
@@ -262,7 +262,7 @@
 by (fast elim!: not_emptyE lepoll_1_is_sing)
 
 lemma AC13_AC1_lemma:
-     "\<forall>B \<in> A. f(B)\<noteq>0 & f(B)<=B & f(B) \<lesssim> 1   
+     "\<forall>B \<in> A. f(B)\<noteq>0 \<and> f(B)<=B \<and> f(B) \<lesssim> 1   
       \<Longrightarrow> (\<lambda>x \<in> A. THE y. f(x)={y}) \<in> (\<Prod>X \<in> A. X)"
 apply (rule lam_type)
 apply (drule bspec, assumption)
--- a/src/ZF/AC/AC16_WO4.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/AC16_WO4.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -16,8 +16,8 @@
 
 lemma lemma1:
      "\<lbrakk>Finite(A); 0<m; m \<in> nat\<rbrakk> 
-      \<Longrightarrow> \<exists>a f. Ord(a) & domain(f) = a &   
-                (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m)"
+      \<Longrightarrow> \<exists>a f. Ord(a) \<and> domain(f) = a \<and>   
+                (\<Union>b<a. f`b) = A \<and> (\<forall>b<a. f`b \<lesssim> m)"
 apply (unfold Finite_def)
 apply (erule bexE)
 apply (drule eqpoll_sym [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]])
@@ -47,7 +47,7 @@
 
 lemmas lepoll_paired = paired_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll]
 
-lemma lemma2: "\<exists>y R. well_ord(y,R) & x \<inter> y = 0 & \<not>y \<lesssim> z & \<not>Finite(y)"
+lemma lemma2: "\<exists>y R. well_ord(y,R) \<and> x \<inter> y = 0 \<and> \<not>y \<lesssim> z \<and> \<not>Finite(y)"
 apply (rule_tac x = "{{a,x}. a \<in> nat \<union> Hartog (z) }" in exI)
 apply (rule well_ord_Un [OF Ord_nat [THEN well_ord_Memrel] 
                          Ord_Hartog [THEN well_ord_Memrel], THEN exE])
@@ -133,7 +133,7 @@
 
 lemma eqpoll_sum_imp_Diff_lepoll_lemma [rule_format]:
      "\<lbrakk>k \<in> nat; m \<in> nat\<rbrakk> 
-      \<Longrightarrow> \<forall>A B. A \<approx> k #+ m & k \<lesssim> B & B \<subseteq> A \<longrightarrow> A-B \<lesssim> m"
+      \<Longrightarrow> \<forall>A B. A \<approx> k #+ m \<and> k \<lesssim> B \<and> B \<subseteq> A \<longrightarrow> A-B \<lesssim> m"
 apply (induct_tac "k")
 apply (simp add: add_0)
 apply (blast intro: eqpoll_imp_lepoll lepoll_trans
@@ -162,7 +162,7 @@
 
 lemma eqpoll_sum_imp_Diff_eqpoll_lemma [rule_format]:
      "\<lbrakk>k \<in> nat; m \<in> nat\<rbrakk> 
-      \<Longrightarrow> \<forall>A B. A \<approx> k #+ m & k \<approx> B & B \<subseteq> A \<longrightarrow> A-B \<approx> m"
+      \<Longrightarrow> \<forall>A B. A \<approx> k #+ m \<and> k \<approx> B \<and> B \<subseteq> A \<longrightarrow> A-B \<approx> m"
 apply (induct_tac "k")
 apply (force dest!: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_0_is_0])
 apply (intro allI impI)
@@ -208,10 +208,10 @@
   defines k_def:     "k   \<equiv> succ(l)"
       and MM_def:    "MM  \<equiv> {v \<in> t_n. succ(k) \<lesssim> v \<inter> y}"
       and LL_def:    "LL  \<equiv> {v \<inter> y. v \<in> MM}"
-      and GG_def:    "GG  \<equiv> \<lambda>v \<in> LL. (THE w. w \<in> MM & v \<subseteq> w) - v"
-      and s_def:     "s(u) \<equiv> {v \<in> t_n. u \<in> v & k \<lesssim> v \<inter> y}"
+      and GG_def:    "GG  \<equiv> \<lambda>v \<in> LL. (THE w. w \<in> MM \<and> v \<subseteq> w) - v"
+      and s_def:     "s(u) \<equiv> {v \<in> t_n. u \<in> v \<and> k \<lesssim> v \<inter> y}"
   assumes all_ex:    "\<forall>z \<in> {z \<in> Pow(x \<union> y) . z \<approx> succ(k)}.
-                       \<exists>! w. w \<in> t_n & z \<subseteq> w "
+                       \<exists>! w. w \<in> t_n \<and> z \<subseteq> w "
     and disjoint[iff]:  "x \<inter> y = 0"
     and "includes":  "t_n \<subseteq> {v \<in> Pow(x \<union> y). v \<approx> succ(k #+ m)}"
     and WO_R[iff]:      "well_ord(y,R)"
@@ -265,7 +265,7 @@
 
 lemma ex1_superset_a:
      "\<lbrakk>l \<approx> a;  a \<subseteq> y;  b \<in> y - a;  u \<in> x\<rbrakk>   
-      \<Longrightarrow> \<exists>! c. c \<in> s(u) & a \<subseteq> c & b \<in> c"
+      \<Longrightarrow> \<exists>! c. c \<in> s(u) \<and> a \<subseteq> c \<and> b \<in> c"
 apply (rule all_ex [simplified k_def, THEN ballE])
  apply (erule ex1E)
  apply (rule_tac a = w in ex1I, blast intro: sI)
@@ -277,7 +277,7 @@
 lemma the_eq_cons:
      "\<lbrakk>\<forall>v \<in> s(u). succ(l) \<approx> v \<inter> y;   
          l \<approx> a;  a \<subseteq> y;  b \<in> y - a;  u \<in> x\<rbrakk>    
-      \<Longrightarrow> (THE c. c \<in> s(u) & a \<subseteq> c & b \<in> c) \<inter> y = cons(b, a)"
+      \<Longrightarrow> (THE c. c \<in> s(u) \<and> a \<subseteq> c \<and> b \<in> c) \<inter> y = cons(b, a)"
 apply (frule ex1_superset_a [THEN theI], assumption+)
 apply (rule set_eq_cons)
 apply (fast+)
@@ -289,7 +289,7 @@
       \<Longrightarrow> y \<lesssim> {v \<in> s(u). a \<subseteq> v}"
 apply (rule Diff_Finite_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll, 
                                 THEN lepoll_trans],  fast+)
-apply (rule_tac  f3 = "\<lambda>b \<in> y-a. THE c. c \<in> s (u) & a \<subseteq> c & b \<in> c" 
+apply (rule_tac  f3 = "\<lambda>b \<in> y-a. THE c. c \<in> s (u) \<and> a \<subseteq> c \<and> b \<in> c" 
         in exI [THEN lepoll_def [THEN def_imp_iff, THEN iffD2]])
 apply (simp add: inj_def)
 apply (rule conjI)
@@ -343,7 +343,7 @@
 
 lemma cons_cons_in:
      "\<lbrakk>z \<in> xa \<inter> (x - {u}); l \<approx> a; a \<subseteq> y; u \<in> x\<rbrakk>   
-      \<Longrightarrow> \<exists>! w. w \<in> t_n & cons(z, cons(u, a)) \<subseteq> w"
+      \<Longrightarrow> \<exists>! w. w \<in> t_n \<and> cons(z, cons(u, a)) \<subseteq> w"
 apply (rule all_ex [THEN bspec])
 apply (unfold k_def)
 apply (fast intro!: cons_eqpoll_succ elim: eqpoll_sym)
@@ -410,7 +410,7 @@
 (* ********************************************************************** *)
 
 lemma unique_superset_in_MM:
-     "v \<in> LL \<Longrightarrow> \<exists>! w. w \<in> MM & v \<subseteq> w"
+     "v \<in> LL \<Longrightarrow> \<exists>! w. w \<in> MM \<and> v \<subseteq> w"
 apply (unfold MM_def LL_def, safe, fast)
 apply (rule lepoll_imp_eqpoll_subset [THEN exE], assumption)
 apply (rule_tac x = x in all_ex [THEN ballE]) 
@@ -430,7 +430,7 @@
 by (unfold LL_def, fast)
 
 lemma in_LL_eq_Int: 
-     "v \<in> LL \<Longrightarrow> v = (THE x. x \<in> MM & v \<subseteq> x) \<inter> y"
+     "v \<in> LL \<Longrightarrow> v = (THE x. x \<in> MM \<and> v \<subseteq> x) \<inter> y"
 apply (unfold LL_def, clarify)
 apply (subst unique_superset_in_MM [THEN the_equality2])
 apply (auto simp add: Int_in_LL)
@@ -440,7 +440,7 @@
 by (erule unique_superset_in_MM [THEN theI, THEN conjunct1]) 
 
 lemma the_in_MM_subset:
-     "v \<in> LL \<Longrightarrow> (THE x. x \<in> MM & v \<subseteq> x) \<subseteq> x \<union> y"
+     "v \<in> LL \<Longrightarrow> (THE x. x \<in> MM \<and> v \<subseteq> x) \<subseteq> x \<union> y"
 apply (drule unique_superset1)
 apply (unfold MM_def)
 apply (fast dest!: unique_superset1 "includes" [THEN subsetD])
@@ -461,7 +461,7 @@
 apply (rule Infinite) 
 done
 
-lemma ex_subset_eqpoll_n: "n \<in> nat \<Longrightarrow> \<exists>z. z \<subseteq> y & n \<approx> z"
+lemma ex_subset_eqpoll_n: "n \<in> nat \<Longrightarrow> \<exists>z. z \<subseteq> y \<and> n \<approx> z"
 apply (erule nat_lepoll_imp_ex_eqpoll_n)
 apply (rule lepoll_trans [OF nat_lepoll_ordertype]) 
 apply (rule ordertype_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll]) 
@@ -543,7 +543,7 @@
 done
 
 lemma "conclusion":
-     "\<exists>a f. Ord(a) & domain(f) = a & (\<Union>b<a. f ` b) = x & (\<forall>b<a. f ` b \<lesssim> m)"
+     "\<exists>a f. Ord(a) \<and> domain(f) = a \<and> (\<Union>b<a. f ` b) = x \<and> (\<forall>b<a. f ` b \<lesssim> m)"
 apply (rule well_ord_LL [THEN exE])
 apply (rename_tac S)
 apply (rule_tac x = "ordertype (LL,S)" in exI)
--- a/src/ZF/AC/AC16_lemmas.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/AC16_lemmas.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -106,7 +106,7 @@
 by fast
 
 lemma Union_in_lemma [rule_format]:
-     "n \<in> nat \<Longrightarrow> \<forall>z. (\<forall>y \<in> z. Ord(y)) & z\<approx>n & z\<noteq>0 \<longrightarrow> \<Union>(z) \<in> z"
+     "n \<in> nat \<Longrightarrow> \<forall>z. (\<forall>y \<in> z. Ord(y)) \<and> z\<approx>n \<and> z\<noteq>0 \<longrightarrow> \<Union>(z) \<in> z"
 apply (induct_tac "n")
 apply (fast dest!: eqpoll_imp_lepoll [THEN lepoll_0_is_0])
 apply (intro allI impI)
@@ -196,7 +196,7 @@
 apply (fast elim!: bij_is_surj [THEN image_vimage_eq])
 done
 
-lemma WO2_imp_ex_Card: "WO2 \<Longrightarrow> \<exists>a. Card(a) & X\<approx>a"
+lemma WO2_imp_ex_Card: "WO2 \<Longrightarrow> \<exists>a. Card(a) \<and> X\<approx>a"
 apply (unfold WO2_def)
 apply (drule spec [of _ X])
 apply (blast intro: Card_cardinal eqpoll_trans
@@ -220,7 +220,7 @@
 apply (blast intro: subsets_eqpoll subsets_eqpoll_X eqpoll_sym eqpoll_trans) 
 done
 
-lemma well_ord_imp_ex_Card: "well_ord(X,R) \<Longrightarrow> \<exists>a. Card(a) & X\<approx>a"
+lemma well_ord_imp_ex_Card: "well_ord(X,R) \<Longrightarrow> \<exists>a. Card(a) \<and> X\<approx>a"
 by (fast elim!: well_ord_cardinal_eqpoll [THEN eqpoll_sym] 
          intro!: Card_cardinal)
 
--- a/src/ZF/AC/AC17_AC1.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/AC17_AC1.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -13,7 +13,7 @@
 
 
 (** AC0 is equivalent to AC1.  
-    AC0 comes from Suppes, AC1 from Rubin & Rubin **)
+    AC0 comes from Suppes, AC1 from Rubin \<and> Rubin **)
 
 lemma AC0_AC1_lemma: "\<lbrakk>f:(\<Prod>X \<in> A. X); D \<subseteq> A\<rbrakk> \<Longrightarrow> \<exists>g. g:(\<Prod>X \<in> D. X)"
 by (fast intro!: lam_type apply_type)
@@ -262,7 +262,7 @@
 
 
 (* ********************************************************************** *)
-(* AC5 \<Longrightarrow> AC4, Rubin & Rubin, p. 11                                      *)
+(* AC5 \<Longrightarrow> AC4, Rubin \<and> Rubin, p. 11                                      *)
 (* ********************************************************************** *)
 
 lemma AC5_AC4_aux1: "R \<subseteq> A*B \<Longrightarrow> (\<lambda>x \<in> R. fst(x)) \<in> R -> A"
--- a/src/ZF/AC/AC18_AC19.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/AC18_AC19.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -56,7 +56,7 @@
 (* ********************************************************************** *)
 
 lemma RepRep_conj: 
-        "\<lbrakk>A \<noteq> 0; 0 \<notin> A\<rbrakk> \<Longrightarrow> {uu(a). a \<in> A} \<noteq> 0 & 0 \<notin> {uu(a). a \<in> A}"
+        "\<lbrakk>A \<noteq> 0; 0 \<notin> A\<rbrakk> \<Longrightarrow> {uu(a). a \<in> A} \<noteq> 0 \<and> 0 \<notin> {uu(a). a \<in> A}"
 apply (unfold uu_def, auto) 
 apply (blast dest!: sym [THEN RepFun_eq_0_iff [THEN iffD1]])
 done
--- a/src/ZF/AC/AC7_AC9.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/AC7_AC9.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -48,7 +48,7 @@
 by (unfold AC6_def AC7_def, blast)
 
 (* ********************************************************************** *)
-(* AC7 \<Longrightarrow> AC6, Rubin & Rubin p. 12, Theorem 2.8                          *)
+(* AC7 \<Longrightarrow> AC6, Rubin \<and> Rubin p. 12, Theorem 2.8                          *)
 (* The case of the empty family of sets added in order to complete        *)
 (* the proof.                                                             *)
 (* ********************************************************************** *)
@@ -86,7 +86,7 @@
 (* ********************************************************************** *)
 
 lemma AC1_AC8_lemma1: 
-        "\<forall>B \<in> A. \<exists>B1 B2. B=<B1,B2> & B1 \<approx> B2   
+        "\<forall>B \<in> A. \<exists>B1 B2. B=<B1,B2> \<and> B1 \<approx> B2   
         \<Longrightarrow> 0 \<notin> { bij(fst(B),snd(B)). B \<in> A }"
 apply (unfold eqpoll_def, auto)
 done
@@ -104,13 +104,13 @@
 
 (* ********************************************************************** *)
 (* AC8 \<Longrightarrow> AC9                                                            *)
-(*  - this proof replaces the following two from Rubin & Rubin:           *)
+(*  - this proof replaces the following two from Rubin \<and> Rubin:           *)
 (*    AC8 \<Longrightarrow> AC1 and AC1 \<Longrightarrow> AC9                                         *)
 (* ********************************************************************** *)
 
 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> & B1 \<approx> B2"
+      \<Longrightarrow> \<forall>B \<in> A*A. \<exists>B1 B2. B=<B1,B2> \<and> B1 \<approx> B2"
 by fast
 
 lemma AC8_AC9: "AC8 \<Longrightarrow> AC9"
@@ -124,7 +124,7 @@
 (* ********************************************************************** *)
 (* AC9 \<Longrightarrow> AC1                                                            *)
 (* The idea of this proof comes from "Equivalents of the Axiom of Choice" *)
-(* by Rubin & Rubin. But (x * y) is not necessarily equipollent to        *)
+(* by Rubin \<and> Rubin. But (x * y) is not necessarily equipollent to        *)
 (* (x * y) \<union> {0} when y is a set of total functions acting from nat to   *)
 (* \<Union>(A) -- therefore we have used the set (y * nat) instead of y.     *)
 (* ********************************************************************** *)
--- a/src/ZF/AC/AC_Equiv.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/AC_Equiv.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -21,21 +21,21 @@
     "WO1 \<equiv> \<forall>A. \<exists>R. well_ord(A,R)"
 
 definition  
-    "WO2 \<equiv> \<forall>A. \<exists>a. Ord(a) & A\<approx>a"
+    "WO2 \<equiv> \<forall>A. \<exists>a. Ord(a) \<and> A\<approx>a"
 
 definition  
-    "WO3 \<equiv> \<forall>A. \<exists>a. Ord(a) & (\<exists>b. b \<subseteq> a & A\<approx>b)"
+    "WO3 \<equiv> \<forall>A. \<exists>a. Ord(a) \<and> (\<exists>b. b \<subseteq> a \<and> A\<approx>b)"
 
 definition  
-    "WO4(m) \<equiv> \<forall>A. \<exists>a f. Ord(a) & domain(f)=a &   
-                         (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m)"
+    "WO4(m) \<equiv> \<forall>A. \<exists>a f. Ord(a) \<and> domain(f)=a \<and>   
+                         (\<Union>b<a. f`b) = A \<and> (\<forall>b<a. f`b \<lesssim> m)"
 
 definition  
-    "WO5 \<equiv> \<exists>m \<in> nat. 1\<le>m & WO4(m)"
+    "WO5 \<equiv> \<exists>m \<in> nat. 1\<le>m \<and> WO4(m)"
 
 definition  
-    "WO6 \<equiv> \<forall>A. \<exists>m \<in> nat. 1\<le>m & (\<exists>a f. Ord(a) & domain(f)=a
-                               & (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m))"
+    "WO6 \<equiv> \<forall>A. \<exists>m \<in> nat. 1\<le>m \<and> (\<exists>a f. Ord(a) \<and> domain(f)=a
+                               \<and> (\<Union>b<a. f`b) = A \<and> (\<forall>b<a. f`b \<lesssim> m))"
 
 definition  
     "WO7 \<equiv> \<forall>A. Finite(A) \<longleftrightarrow> (\<forall>R. well_ord(A,R) \<longrightarrow> well_ord(A,converse(R)))"
@@ -51,7 +51,7 @@
 
 definition
   sets_of_size_between :: "[i, i, i] => o"  where
-    "sets_of_size_between(A,m,n) \<equiv> \<forall>B \<in> A. m \<lesssim> B & B \<lesssim> n"
+    "sets_of_size_between(A,m,n) \<equiv> \<forall>B \<in> A. m \<lesssim> B \<and> B \<lesssim> n"
 
 
 (* Axioms of Choice *)  
@@ -62,7 +62,7 @@
     "AC1 \<equiv> \<forall>A. 0\<notin>A \<longrightarrow> (\<exists>f. f \<in> (\<Prod>X \<in> A. X))"
 
 definition
-    "AC2 \<equiv> \<forall>A. 0\<notin>A & pairwise_disjoint(A)   
+    "AC2 \<equiv> \<forall>A. 0\<notin>A \<and> pairwise_disjoint(A)   
                    \<longrightarrow> (\<exists>C. \<forall>B \<in> A. \<exists>y. B \<inter> C = {y})"
 definition
     "AC3 \<equiv> \<forall>A B. \<forall>f \<in> A->B. \<exists>g. g \<in> (\<Prod>x \<in> {a \<in> A. f`a\<noteq>0}. f`x)"
@@ -77,10 +77,10 @@
     "AC6 \<equiv> \<forall>A. 0\<notin>A \<longrightarrow> (\<Prod>B \<in> A. B)\<noteq>0"
 
 definition
-    "AC7 \<equiv> \<forall>A. 0\<notin>A & (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) \<longrightarrow> (\<Prod>B \<in> A. B) \<noteq> 0"
+    "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> & B1\<approx>B2)   
+    "AC8 \<equiv> \<forall>A. (\<forall>B \<in> A. \<exists>B1 B2. B=<B1,B2> \<and> B1\<approx>B2)   
                    \<longrightarrow> (\<exists>f. \<forall>B \<in> A. f`B \<in> bij(fst(B),snd(B)))"
 
 definition
@@ -89,45 +89,45 @@
 
 definition
     "AC10(n) \<equiv>  \<forall>A. (\<forall>B \<in> A. \<not>Finite(B)) \<longrightarrow>   
-                   (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &   
-                   sets_of_size_between(f`B, 2, succ(n)) & \<Union>(f`B)=B))"
+                   (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) \<and>   
+                   sets_of_size_between(f`B, 2, succ(n)) \<and> \<Union>(f`B)=B))"
 
 definition
-    "AC11 \<equiv> \<exists>n \<in> nat. 1\<le>n & AC10(n)"
+    "AC11 \<equiv> \<exists>n \<in> nat. 1\<le>n \<and> AC10(n)"
 
 definition
     "AC12 \<equiv> \<forall>A. (\<forall>B \<in> A. \<not>Finite(B)) \<longrightarrow>
-                 (\<exists>n \<in> nat. 1\<le>n & (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &   
-                      sets_of_size_between(f`B, 2, succ(n)) & \<Union>(f`B)=B)))"
+                 (\<exists>n \<in> nat. 1\<le>n \<and> (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) \<and>   
+                      sets_of_size_between(f`B, 2, succ(n)) \<and> \<Union>(f`B)=B)))"
 
 definition
-    "AC13(m) \<equiv> \<forall>A. 0\<notin>A \<longrightarrow> (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 & f`B \<subseteq> B & f`B \<lesssim> m)"
+    "AC13(m) \<equiv> \<forall>A. 0\<notin>A \<longrightarrow> (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 \<and> f`B \<subseteq> B \<and> f`B \<lesssim> m)"
 
 definition
-    "AC14 \<equiv> \<exists>m \<in> nat. 1\<le>m & AC13(m)"
+    "AC14 \<equiv> \<exists>m \<in> nat. 1\<le>m \<and> AC13(m)"
 
 definition
     "AC15 \<equiv> \<forall>A. 0\<notin>A \<longrightarrow> 
-                 (\<exists>m \<in> nat. 1\<le>m & (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 & f`B \<subseteq> B & f`B \<lesssim> m))"
+                 (\<exists>m \<in> nat. 1\<le>m \<and> (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 \<and> f`B \<subseteq> B \<and> f`B \<lesssim> m))"
 
 definition
     "AC16(n, k)  \<equiv> 
        \<forall>A. \<not>Finite(A) \<longrightarrow>   
-           (\<exists>T. T \<subseteq> {X \<in> Pow(A). X\<approx>succ(n)} &   
-           (\<forall>X \<in> {X \<in> Pow(A). X\<approx>succ(k)}. \<exists>! Y. Y \<in> T & X \<subseteq> Y))"
+           (\<exists>T. T \<subseteq> {X \<in> Pow(A). X\<approx>succ(n)} \<and>   
+           (\<forall>X \<in> {X \<in> Pow(A). X\<approx>succ(k)}. \<exists>! Y. Y \<in> T \<and> X \<subseteq> Y))"
 
 definition
     "AC17 \<equiv> \<forall>A. \<forall>g \<in> (Pow(A)-{0} -> A) -> Pow(A)-{0}.   
                    \<exists>f \<in> Pow(A)-{0} -> A. f`(g`f) \<in> g`f"
 
 locale AC18 =
-  assumes AC18: "A\<noteq>0 & (\<forall>a \<in> A. B(a) \<noteq> 0) \<longrightarrow>
+  assumes AC18: "A\<noteq>0 \<and> (\<forall>a \<in> A. B(a) \<noteq> 0) \<longrightarrow>
     ((\<Inter>a \<in> A. \<Union>b \<in> B(a). X(a,b)) =   
       (\<Union>f \<in> \<Prod>a \<in> A. B(a). \<Inter>a \<in> A. X(a, f`a)))"
   \<comment> \<open>AC18 cannot be expressed within the object-logic\<close>
 
 definition
-    "AC19 \<equiv> \<forall>A. A\<noteq>0 & 0\<notin>A \<longrightarrow> ((\<Inter>a \<in> A. \<Union>b \<in> a. b) =   
+    "AC19 \<equiv> \<forall>A. A\<noteq>0 \<and> 0\<notin>A \<longrightarrow> ((\<Inter>a \<in> A. \<Union>b \<in> a. b) =   
                    (\<Union>f \<in> (\<Prod>B \<in> A. B). \<Inter>a \<in> A. f`a))"
 
 
@@ -196,7 +196,7 @@
      "\<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 & f`<x,y> = i" 
+apply (rule_tac x = "\<lambda>x \<in> domain(u). \<mu> i. \<exists>y. <x,y> \<in> u \<and> f`<x,y> = i" 
        in exI)
 apply (rule_tac d = "%y. fst (converse(f) ` y) " in lam_injective)
 apply (fast intro: LeastI2 nat_into_Ord [THEN Ord_in_Ord] 
--- a/src/ZF/AC/Cardinal_aux.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/Cardinal_aux.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -20,14 +20,14 @@
 
 (* j=|A| *)
 lemma lepoll_imp_ex_le_eqpoll:
-     "\<lbrakk>A \<lesssim> i; Ord(i)\<rbrakk> \<Longrightarrow> \<exists>j. j \<le> i & A \<approx> j"
+     "\<lbrakk>A \<lesssim> i; Ord(i)\<rbrakk> \<Longrightarrow> \<exists>j. j \<le> i \<and> A \<approx> j"
 by (blast intro!: lepoll_cardinal_le well_ord_Memrel
                   well_ord_cardinal_eqpoll [THEN eqpoll_sym]
           dest: lepoll_well_ord)
 
 (* j=|A| *)
 lemma lesspoll_imp_ex_lt_eqpoll:
-     "\<lbrakk>A \<prec> i; Ord(i)\<rbrakk> \<Longrightarrow> \<exists>j. j<i & A \<approx> j"
+     "\<lbrakk>A \<prec> i; Ord(i)\<rbrakk> \<Longrightarrow> \<exists>j. j<i \<and> A \<approx> j"
 by (unfold lesspoll_def, blast dest!: lepoll_imp_ex_le_eqpoll elim!: leE)
 
 lemma Un_eqpoll_Inf_Ord:
@@ -65,7 +65,7 @@
 lemma paired_eqpoll: "{{y,z}. y \<in> x} \<approx> x"
 by (unfold eqpoll_def, fast intro!: paired_bij)
 
-lemma ex_eqpoll_disjoint: "\<exists>B. B \<approx> A & B \<inter> C = 0"
+lemma ex_eqpoll_disjoint: "\<exists>B. B \<approx> A \<and> B \<inter> C = 0"
 by (fast intro!: paired_eqpoll equals0I elim: mem_asym)
 
 (*Finally we reach this result.  Surely there's a simpler proof?*)
@@ -108,7 +108,7 @@
 
 lemma UN_fun_lepoll_lemma [rule_format]:
      "\<lbrakk>well_ord(T, R); \<not>Finite(a); Ord(a); n \<in> nat\<rbrakk>
-      \<Longrightarrow> \<forall>f. (\<forall>b \<in> a. f`b \<lesssim> n & f`b \<subseteq> T) \<longrightarrow> (\<Union>b \<in> a. f`b) \<lesssim> a"
+      \<Longrightarrow> \<forall>f. (\<forall>b \<in> a. f`b \<lesssim> n \<and> f`b \<subseteq> T) \<longrightarrow> (\<Union>b \<in> a. f`b) \<lesssim> a"
 apply (induct_tac "n")
 apply (rule allI)
 apply (rule impI)
@@ -126,12 +126,12 @@
 done
 
 lemma UN_fun_lepoll:
-     "\<lbrakk>\<forall>b \<in> a. f`b \<lesssim> n & f`b \<subseteq> T; well_ord(T, R);
+     "\<lbrakk>\<forall>b \<in> a. f`b \<lesssim> n \<and> f`b \<subseteq> T; well_ord(T, R);
          \<not>Finite(a); Ord(a); n \<in> nat\<rbrakk> \<Longrightarrow> (\<Union>b \<in> a. f`b) \<lesssim> a"
 by (blast intro: UN_fun_lepoll_lemma)
 
 lemma UN_lepoll:
-     "\<lbrakk>\<forall>b \<in> a. F(b) \<lesssim> n & F(b) \<subseteq> T; well_ord(T, R);
+     "\<lbrakk>\<forall>b \<in> a. F(b) \<lesssim> n \<and> F(b) \<subseteq> T; well_ord(T, R);
          \<not>Finite(a); Ord(a); n \<in> nat\<rbrakk>
       \<Longrightarrow> (\<Union>b \<in> a. F(b)) \<lesssim> a"
 apply (rule rev_mp)
@@ -153,7 +153,7 @@
 done
 
 lemma lepoll_imp_eqpoll_subset:
-     "a \<lesssim> X \<Longrightarrow> \<exists>Y. Y \<subseteq> X & a \<approx> Y"
+     "a \<lesssim> X \<Longrightarrow> \<exists>Y. Y \<subseteq> X \<and> a \<approx> Y"
 apply (unfold lepoll_def eqpoll_def, clarify)
 apply (blast intro: restrict_bij
              dest: inj_is_fun [THEN fun_is_rel, THEN image_subset])
--- a/src/ZF/AC/DC.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/DC.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -81,13 +81,13 @@
 
 definition
   DC  :: "i => o"  where
-    "DC(a) \<equiv> \<forall>X R. R \<subseteq> Pow(X)*X  &
+    "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)) 
                     \<longrightarrow> (\<exists>f \<in> a->X. \<forall>b<a. <f``b,f`b> \<in> R)"
 
 definition
   DC0 :: o  where
-    "DC0 \<equiv> \<forall>A B R. R \<subseteq> A*B & R\<noteq>0 & range(R) \<subseteq> domain(R) 
+    "DC0 \<equiv> \<forall>A B R. R \<subseteq> A*B \<and> R\<noteq>0 \<and> range(R) \<subseteq> domain(R) 
                     \<longrightarrow> (\<exists>f \<in> nat->domain(R). \<forall>n \<in> nat. <f`n,f`succ(n)>:R)"
 
 definition
@@ -103,7 +103,7 @@
 
   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))  
-                                       & restrict(z2, domain(z1)) = z1}"
+                                       \<and> restrict(z2, domain(z1)) = z1}"
 begin
 
 (* ********************************************************************** *)
@@ -116,7 +116,7 @@
 (* Define XX and RR as follows:                                           *)
 (*                                                                        *)
 (*       XX = (\<Union>n \<in> nat. {f \<in> n->X. \<forall>k \<in> n. <f``k, f`k> \<in> R})           *)
-(*       f RR g iff domain(g)=succ(domain(f)) &                           *)
+(*       f RR g iff domain(g)=succ(domain(f)) \<and>                           *)
 (*              restrict(g, domain(f)) = f                                *)
 (*                                                                        *)
 (* Then RR satisfies the hypotheses of DC.                                *)
@@ -170,8 +170,8 @@
 
 lemma lemma2:
      "\<lbrakk>\<forall>n \<in> nat. <f`n, f`succ(n)> \<in> RR;  f \<in> nat -> XX;  n \<in> nat\<rbrakk>   
-      \<Longrightarrow> \<exists>k \<in> nat. f`succ(n) \<in> k -> X & n \<in> k   
-                  & <f`succ(n)``n, f`succ(n)`n> \<in> R"
+      \<Longrightarrow> \<exists>k \<in> nat. f`succ(n) \<in> k -> X \<and> n \<in> k   
+                  \<and> <f`succ(n)``n, f`succ(n)`n> \<in> R"
 apply (induct_tac "n")
 apply (drule apply_type [OF _ nat_1I])
 apply (drule bspec [OF _ nat_0I])
@@ -267,10 +267,10 @@
     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. 
-           (domain(z2)=succ(\<Union>f \<in> z1. domain(f)) &
+           (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)) &
-                        (\<forall>f \<in> z1. restrict(g, 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>})}                                          
                                                                           
    Then XX and RR satisfy the hypotheses of DC(omega).                    
@@ -302,20 +302,20 @@
       and RR_def:
          "RR \<equiv> {<z1,z2>:Fin(XX)*XX. 
                   (domain(z2)=succ(\<Union>f \<in> z1. domain(f))  
-                    & (\<forall>f \<in> z1. restrict(z2, domain(f)) = 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))  
-                     & (\<forall>f \<in> z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}"
+                     \<and> (\<forall>f \<in> z1. restrict(g, domain(f)) = f)) \<and> z2={<0,x>})}"
       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))
-                                    & (\<Union>f \<in> z1. domain(f)) = b  
-                                    & (\<forall>f \<in> z1. restrict(z2,domain(f)) = f))}"
+                                    \<and> (\<Union>f \<in> z1. domain(f)) = b  
+                                    \<and> (\<forall>f \<in> z1. restrict(z2,domain(f)) = f))}"
 begin
 
 lemma lemma4:
      "\<lbrakk>range(R) \<subseteq> domain(R);  x \<in> domain(R)\<rbrakk>   
-      \<Longrightarrow> RR \<subseteq> Pow(XX)*XX &   
+      \<Longrightarrow> RR \<subseteq> Pow(XX)*XX \<and>   
              (\<forall>Y \<in> Pow(XX). Y \<prec> nat \<longrightarrow> (\<exists>x \<in> XX. <Y,x>:RR))"
 apply (rule conjI)
 apply (force dest!: FinD [THEN PowI] simp add: RR_def)
@@ -323,7 +323,7 @@
 apply (drule Finite_Fin [OF lesspoll_nat_is_Finite PowD], assumption)
 apply (case_tac
        "\<exists>g \<in> XX. domain (g) =
-             succ(\<Union>f \<in> Y. domain(f)) & (\<forall>f\<in>Y. restrict(g, domain(f)) = f)")
+             succ(\<Union>f \<in> Y. domain(f)) \<and> (\<forall>f\<in>Y. restrict(g, domain(f)) = f)")
 apply (simp add: RR_def, blast)
 apply (safe del: domainE)
 apply (unfold XX_def RR_def)
@@ -436,7 +436,7 @@
 
 lemma lemma2: 
      "\<lbrakk>allRR; f \<in> nat->XX; range(R) \<subseteq> domain(R); x \<in> domain(R); n \<in> nat\<rbrakk>
-      \<Longrightarrow> f`n \<in> succ(n) -> domain(R) & (\<forall>i \<in> n. <f`n`i, f`n`succ(i)>:R)"
+      \<Longrightarrow> f`n \<in> succ(n) -> domain(R) \<and> (\<forall>i \<in> n. <f`n`i, f`n`succ(i)>:R)"
 apply (unfold allRR_def)
 apply (drule ospec)
 apply (erule ltI [OF _ Ord_nat])
@@ -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) & z2 \<notin> z1}" 
+apply (erule_tac x = "{<z1,z2> \<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])
--- a/src/ZF/AC/HH.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/HH.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -32,7 +32,7 @@
      "B \<subseteq> A \<Longrightarrow> X-(\<Union>a \<in> A. P(a)) = X-(\<Union>a \<in> A-B. P(a))-(\<Union>b \<in> B. P(b))"
 by fast
 
-lemma Ord_DiffE: "\<lbrakk>c \<in> a-b; b<a\<rbrakk> \<Longrightarrow> c=b | b<c & c<a"
+lemma Ord_DiffE: "\<lbrakk>c \<in> a-b; b<a\<rbrakk> \<Longrightarrow> c=b | b<c \<and> c<a"
 apply (erule ltE)
 apply (drule Ord_linear [of _ c])
 apply (fast elim: Ord_in_Ord)
--- a/src/ZF/AC/Hartog.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/Hartog.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -19,7 +19,7 @@
 
 lemma Ord_lepoll_imp_ex_well_ord:
      "\<lbrakk>Ord(a); a \<lesssim> X\<rbrakk> 
-      \<Longrightarrow> \<exists>Y. Y \<subseteq> X & (\<exists>R. well_ord(Y,R) & ordertype(Y,R)=a)"
+      \<Longrightarrow> \<exists>Y. Y \<subseteq> X \<and> (\<exists>R. well_ord(Y,R) \<and> ordertype(Y,R)=a)"
 apply (unfold lepoll_def)
 apply (erule exE)
 apply (intro exI conjI)
@@ -36,7 +36,7 @@
 done
 
 lemma Ord_lepoll_imp_eq_ordertype:
-     "\<lbrakk>Ord(a); a \<lesssim> X\<rbrakk> \<Longrightarrow> \<exists>Y. Y \<subseteq> X & (\<exists>R. R \<subseteq> X*X & ordertype(Y,R)=a)"
+     "\<lbrakk>Ord(a); a \<lesssim> X\<rbrakk> \<Longrightarrow> \<exists>Y. Y \<subseteq> X \<and> (\<exists>R. R \<subseteq> X*X \<and> ordertype(Y,R)=a)"
 apply (drule Ord_lepoll_imp_ex_well_ord, assumption, clarify)
 apply (intro exI conjI)
 apply (erule_tac [3] ordertype_Int, auto) 
@@ -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> & ordertype(Y,R)=b}"
+        a \<in> {b. Z \<in> Pow(X)*Pow(X*X), \<exists>Y R. Z=<Y,R> \<and> ordertype(Y,R)=b}"
 apply (intro allI impI)
 apply (elim allE impE, assumption)
 apply (blast dest!: Ord_lepoll_imp_eq_ordertype intro: sym) 
@@ -54,7 +54,7 @@
 lemma Ords_lepoll_set: "\<forall>a. Ord(a) \<longrightarrow> a \<lesssim> X \<Longrightarrow> P"
 by (erule Ords_lepoll_set_lemma [THEN Ords_in_set])
 
-lemma ex_Ord_not_lepoll: "\<exists>a. Ord(a) & \<not>a \<lesssim> X"
+lemma ex_Ord_not_lepoll: "\<exists>a. Ord(a) \<and> \<not>a \<lesssim> X"
 apply (rule ccontr)
 apply (best intro: Ords_lepoll_set) 
 done
--- a/src/ZF/AC/WO1_AC.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/WO1_AC.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -89,8 +89,8 @@
 
 lemma lemma2:
      "\<lbrakk>WO1; \<not>Finite(B); 1\<le>n\<rbrakk>   
-      \<Longrightarrow> \<exists>C \<in> Pow(Pow(B)). pairwise_disjoint(C) &   
-                sets_of_size_between(C, 2, succ(n)) &   
+      \<Longrightarrow> \<exists>C \<in> Pow(Pow(B)). pairwise_disjoint(C) \<and>   
+                sets_of_size_between(C, 2, succ(n)) \<and>   
                 \<Union>(C)=B"
 apply (drule lemma2_1 [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], 
        assumption)
--- a/src/ZF/AC/WO1_WO7.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/WO1_WO7.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -2,7 +2,7 @@
     Author:     Lawrence C Paulson, CU Computer Laboratory
     Copyright   1998  University of Cambridge
 
-WO7 \<longleftrightarrow> LEMMA \<longleftrightarrow> WO1 (Rubin & Rubin p. 5)
+WO7 \<longleftrightarrow> LEMMA \<longleftrightarrow> WO1 (Rubin \<and> Rubin p. 5)
 LEMMA is the sentence denoted by (**)
 
 Also, WO1 \<longleftrightarrow> WO8
@@ -14,7 +14,7 @@
 
 definition
     "LEMMA \<equiv>
-     \<forall>X. \<not>Finite(X) \<longrightarrow> (\<exists>R. well_ord(X,R) & \<not>well_ord(X,converse(R)))"
+     \<forall>X. \<not>Finite(X) \<longrightarrow> (\<exists>R. well_ord(X,R) \<and> \<not>well_ord(X,converse(R)))"
 
 (* ********************************************************************** *)
 (* It is easy to see that WO7 is equivalent to (**)                       *)
@@ -93,14 +93,14 @@
 
 
 (* ********************************************************************** *)
-(*            The proof of WO8 \<longleftrightarrow> WO1 (Rubin & Rubin p. 6)               *)
+(*            The proof of WO8 \<longleftrightarrow> WO1 (Rubin \<and> Rubin p. 6)               *)
 (* ********************************************************************** *)
 
 lemma WO1_WO8: "WO1 \<Longrightarrow> WO8"
 by (unfold WO1_def WO8_def, fast)
 
 
-(* The implication "WO8 \<Longrightarrow> WO1": a faithful image of Rubin & Rubin's proof*)
+(* The implication "WO8 \<Longrightarrow> WO1": a faithful image of Rubin \<and> Rubin's proof*)
 lemma WO8_WO1: "WO8 \<Longrightarrow> WO1"
 apply (unfold WO1_def WO8_def)
 apply (rule allI)
--- a/src/ZF/AC/WO2_AC16.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/AC/WO2_AC16.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -9,7 +9,7 @@
 The first instance is trivial, the third not difficult, but the second
 is very complicated requiring many lemmas.
 We also need to prove that at any stage gamma the set 
-(s - \<Union>(...) - k_gamma)   (Rubin & Rubin page 15)
+(s - \<Union>(...) - k_gamma)   (Rubin \<and> Rubin page 15)
 contains m distinct elements (in fact is equipollent to s)
 *)
 
@@ -22,7 +22,7 @@
     "recfunAC16(f,h,i,a) \<equiv> 
          transrec2(i, 0, 
               %g r. if (\<exists>y \<in> r. h`g \<subseteq> y) then r
-                    else r \<union> {f`(\<mu> i. h`g \<subseteq> f`i & 
+                    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))))})"
 
 (* ********************************************************************** *)
@@ -36,7 +36,7 @@
      "recfunAC16(f,h,succ(i),a) =   
       (if (\<exists>y \<in> recfunAC16(f,h,i,a). h ` i \<subseteq> y) then recfunAC16(f,h,i,a)  
        else recfunAC16(f,h,i,a) \<union>  
-            {f ` (\<mu> j. h ` i \<subseteq> f ` j &   
+            {f ` (\<mu> j. h ` i \<subseteq> f ` j \<and>   
              (\<forall>b<a. (h`b \<subseteq> f`j   
               \<longrightarrow> (\<forall>t \<in> recfunAC16(f,h,i,a). \<not> h`b \<subseteq> t))))})"
 apply (simp add: recfunAC16_def)
@@ -86,7 +86,7 @@
 
 
 lemma lemma3_1:
-    "\<lbrakk>\<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) \<longrightarrow> (\<exists>! Y. Y \<in> F(y) & f(z)<=Y);
+    "\<lbrakk>\<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) \<longrightarrow> (\<exists>! Y. Y \<in> F(y) \<and> f(z)<=Y);
         \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j); j\<le>i; i<x; z<a;
         V \<in> F(i); f(z)<=V; W \<in> F(j); f(z)<=W\<rbrakk>   
      \<Longrightarrow> V = W"
@@ -96,7 +96,7 @@
 
 
 lemma lemma3:
-    "\<lbrakk>\<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) \<longrightarrow> (\<exists>! Y. Y \<in> F(y) & f(z)<=Y);
+    "\<lbrakk>\<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) \<longrightarrow> (\<exists>! Y. Y \<in> F(y) \<and> f(z)<=Y);
         \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j); i<x; j<x; z<a;   
         V \<in> F(i); f(z)<=V; W \<in> F(j); f(z)<=W\<rbrakk>   
      \<Longrightarrow> V = W"
@@ -106,25 +106,25 @@
 done
 
 lemma lemma4:
-     "\<lbrakk>\<forall>y<x. F(y) \<subseteq> X &   
+     "\<lbrakk>\<forall>y<x. F(y) \<subseteq> X \<and>   
                 (\<forall>x<a. x < y | (\<exists>Y \<in> F(y). h(x) \<subseteq> Y) \<longrightarrow>   
-                 (\<exists>! Y. Y \<in> F(y) & h(x) \<subseteq> Y)); 
+                 (\<exists>! Y. Y \<in> F(y) \<and> h(x) \<subseteq> Y)); 
          x < a\<rbrakk>   
       \<Longrightarrow> \<forall>y<x. \<forall>z<a. z < y | (\<exists>Y \<in> F(y). h(z) \<subseteq> Y) \<longrightarrow>   
-                       (\<exists>! Y. Y \<in> F(y) & h(z) \<subseteq> Y)"
+                       (\<exists>! Y. Y \<in> F(y) \<and> h(z) \<subseteq> Y)"
 apply (intro oallI impI)
 apply (drule ospec, assumption, clarify)
 apply (blast elim!: oallE ) 
 done
 
 lemma lemma5:
-     "\<lbrakk>\<forall>y<x. F(y) \<subseteq> X &   
+     "\<lbrakk>\<forall>y<x. F(y) \<subseteq> X \<and>   
                (\<forall>x<a. x < y | (\<exists>Y \<in> F(y). h(x) \<subseteq> Y) \<longrightarrow>   
-                (\<exists>! Y. Y \<in> F(y) & h(x) \<subseteq> Y));  
+                (\<exists>! Y. Y \<in> F(y) \<and> h(x) \<subseteq> Y));  
          x < a; Limit(x); \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j)\<rbrakk>   
-      \<Longrightarrow> (\<Union>x<x. F(x)) \<subseteq> X &   
+      \<Longrightarrow> (\<Union>x<x. F(x)) \<subseteq> X \<and>   
           (\<forall>xa<a. xa < x | (\<exists>x \<in> \<Union>x<x. F(x). h(xa) \<subseteq> x)   
-                \<longrightarrow> (\<exists>! Y. Y \<in> (\<Union>x<x. F(x)) & h(xa) \<subseteq> Y))"
+                \<longrightarrow> (\<exists>! Y. Y \<in> (\<Union>x<x. F(x)) \<and> h(xa) \<subseteq> Y))"
 apply (rule conjI)
 apply (rule subsetI)
 apply (erule OUN_E)
@@ -148,7 +148,7 @@
   First quite complicated proof of the fact used in the recursive construction
   of the family T_gamma (WO2 \<Longrightarrow> AC16(k #+ m, k)) - the fact that at any stage
   gamma the set (s - \<Union>(...) - k_gamma) is equipollent to s
-  (Rubin & Rubin page 15).
+  (Rubin \<and> Rubin page 15).
 *)
 
 (* ********************************************************************** *)
@@ -177,7 +177,7 @@
 done
 
 lemma Union_lesspoll:
-     "\<lbrakk>\<forall>x \<in> X. x \<lesssim> n & x \<subseteq> T; well_ord(T, R); X \<lesssim> b;   
+     "\<lbrakk>\<forall>x \<in> X. x \<lesssim> n \<and> x \<subseteq> T; well_ord(T, R); X \<lesssim> b;   
          b<a; \<not>Finite(a); Card(a); n \<in> nat\<rbrakk>   
       \<Longrightarrow> \<Union>(X)\<prec>a"
 apply (case_tac "Finite (X)")
@@ -238,7 +238,7 @@
 lemma Least_eq_imp_ex:
      "\<lbrakk>(\<mu> i. w \<in> F(i)) = (\<mu> i. z \<in> F(i));   
          w \<in> (\<Union>i<a. F(i)); z \<in> (\<Union>i<a. F(i))\<rbrakk>   
-      \<Longrightarrow> \<exists>b<a. w \<in> (F(b) - (\<Union>c<b. F(c))) & z \<in> (F(b) - (\<Union>c<b. F(c)))"
+      \<Longrightarrow> \<exists>b<a. w \<in> (F(b) - (\<Union>c<b. F(c))) \<and> z \<in> (F(b) - (\<Union>c<b. F(c)))"
 apply (elim OUN_E)
 apply (drule in_Least_Diff, erule lt_Ord) 
 apply (frule in_Least_Diff, erule lt_Ord) 
@@ -330,8 +330,8 @@
 (* ********************************************************************** *)
 
 lemma lemma6:
-     "\<lbrakk>\<forall>y<succ(j). F(y)<=X & (\<forall>x<a. x<y | P(x,y) \<longrightarrow> Q(x,y)); succ(j)<a\<rbrakk>
-      \<Longrightarrow> F(j)<=X & (\<forall>x<a. x<j | P(x,j) \<longrightarrow> Q(x,j))"
+     "\<lbrakk>\<forall>y<succ(j). F(y)<=X \<and> (\<forall>x<a. x<y | P(x,y) \<longrightarrow> Q(x,y)); succ(j)<a\<rbrakk>
+      \<Longrightarrow> F(j)<=X \<and> (\<forall>x<a. x<j | P(x,j) \<longrightarrow> Q(x,j))"
 by (blast intro!: lt_Ord succI1 [THEN ltI, THEN lt_Ord, THEN le_refl]) 
 
 
@@ -369,7 +369,7 @@
 
 
 lemma subset_imp_eq_lemma:
-     "m \<in> nat \<Longrightarrow> \<forall>A B. A \<subseteq> B & m \<lesssim> A & B \<lesssim> m \<longrightarrow> A=B"
+     "m \<in> nat \<Longrightarrow> \<forall>A B. A \<subseteq> B \<and> m \<lesssim> A \<and> B \<lesssim> m \<longrightarrow> A=B"
 apply (induct_tac "m")
 apply (fast dest!: lepoll_0_is_0)
 apply (intro allI impI)
@@ -406,7 +406,7 @@
          k \<in> nat; m \<in> nat; y<a;   
          h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)});   
          \<not> (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y)\<rbrakk>   
-      \<Longrightarrow> \<exists>X \<in> {Y \<in> Pow(A). Y\<approx>succ(k #+ m)}. h`y \<subseteq> X &   
+      \<Longrightarrow> \<exists>X \<in> {Y \<in> Pow(A). Y\<approx>succ(k #+ m)}. h`y \<subseteq> X \<and>   
                 (\<forall>b<a. h`b \<subseteq> X \<longrightarrow>   
                 (\<forall>T \<in> recfunAC16(f, h, y, a). \<not> h`b \<subseteq> T))"
 apply (erule_tac m1=m in dbl_Diff_eqpoll [THEN ex_subset_eqpoll, THEN bexE], 
@@ -433,7 +433,7 @@
          h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)});   
          f \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k #+ m)});   
          \<not> (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y)\<rbrakk>   
-      \<Longrightarrow> \<exists>c<a. h`y \<subseteq> f`c &   
+      \<Longrightarrow> \<exists>c<a. h`y \<subseteq> f`c \<and>   
                 (\<forall>b<a. h`b \<subseteq> f`c \<longrightarrow>   
                 (\<forall>T \<in> recfunAC16(f, h, y, a). \<not> h`b \<subseteq> T))"
 apply (drule ex_next_set, assumption+)
@@ -451,11 +451,11 @@
 
 lemma lemma8:
      "\<lbrakk>\<forall>x<a. x<j | (\<exists>xa \<in> F(j). P(x, xa))   
-         \<longrightarrow> (\<exists>! Y. Y \<in> F(j) & P(x, Y)); F(j) \<subseteq> X;   
-         L \<in> X; P(j, L) & (\<forall>x<a. P(x, L) \<longrightarrow> (\<forall>xa \<in> F(j). \<not>P(x, xa)))\<rbrakk>   
-      \<Longrightarrow> F(j) \<union> {L} \<subseteq> X &   
+         \<longrightarrow> (\<exists>! Y. Y \<in> F(j) \<and> P(x, Y)); F(j) \<subseteq> X;   
+         L \<in> X; P(j, L) \<and> (\<forall>x<a. P(x, L) \<longrightarrow> (\<forall>xa \<in> F(j). \<not>P(x, xa)))\<rbrakk>   
+      \<Longrightarrow> F(j) \<union> {L} \<subseteq> X \<and>   
           (\<forall>x<a. x\<le>j | (\<exists>xa \<in> (F(j) \<union> {L}). P(x, xa)) \<longrightarrow>   
-                 (\<exists>! Y. Y \<in> (F(j) \<union> {L}) & P(x, Y)))"
+                 (\<exists>! Y. Y \<in> (F(j) \<union> {L}) \<and> P(x, Y)))"
 apply (rule conjI)
 apply (fast intro!: singleton_subsetI)
 apply (rule oallI)
@@ -471,9 +471,9 @@
      "\<lbrakk>b < a; f \<in> bij(a, {Y \<in> Pow(A) . Y\<approx>succ(k #+ m)});   
          h \<in> bij(a, {Y \<in> Pow(A) . Y\<approx>succ(k)});   
          \<not>Finite(a); Card(a); A\<approx>a; k \<in> nat; m \<in> nat\<rbrakk>   
-      \<Longrightarrow> recfunAC16(f, h, b, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)} &   
+      \<Longrightarrow> recfunAC16(f, h, b, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)} \<and>   
           (\<forall>x<a. x < b | (\<exists>Y \<in> recfunAC16(f, h, b, a). h ` x \<subseteq> Y) \<longrightarrow>   
-          (\<exists>! Y. Y \<in> recfunAC16(f, h, b, a) & h ` x \<subseteq> Y))"
+          (\<exists>! Y. Y \<in> recfunAC16(f, h, b, a) \<and> h ` x \<subseteq> Y))"
 apply (erule lt_induct)
 apply (frule lt_Ord)
 apply (erule Ord_cases)
@@ -507,12 +507,12 @@
 (* ********************************************************************** *)
 
 lemma lemma_simp_induct:
-     "\<lbrakk>\<forall>b. b<a \<longrightarrow> F(b) \<subseteq> S & (\<forall>x<a. (x<b | (\<exists>Y \<in> F(b). f`x \<subseteq> Y))  
-                                   \<longrightarrow> (\<exists>! Y. Y \<in> F(b) & f`x \<subseteq> Y));
+     "\<lbrakk>\<forall>b. b<a \<longrightarrow> F(b) \<subseteq> S \<and> (\<forall>x<a. (x<b | (\<exists>Y \<in> F(b). f`x \<subseteq> Y))  
+                                   \<longrightarrow> (\<exists>! Y. Y \<in> F(b) \<and> f`x \<subseteq> Y));
          f \<in> a->f``(a); Limit(a); 
          \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j)\<rbrakk>   
-        \<Longrightarrow> (\<Union>j<a. F(j)) \<subseteq> S &   
-            (\<forall>x \<in> f``a. \<exists>! Y. Y \<in> (\<Union>j<a. F(j)) & x \<subseteq> Y)"
+        \<Longrightarrow> (\<Union>j<a. F(j)) \<subseteq> S \<and>   
+            (\<forall>x \<in> f``a. \<exists>! Y. Y \<in> (\<Union>j<a. F(j)) \<and> x \<subseteq> Y)"
 apply (rule conjI)
 apply (rule subsetI)
 apply (erule OUN_E, blast) 
@@ -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 & (\<forall>x \<in> z. Z(x))"  for Y Z
+apply (rule_tac P="%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 16:51:35 2022 +0100
+++ b/src/ZF/AC/WO6_WO1.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -5,7 +5,7 @@
 The only hard one is WO6 \<Longrightarrow> WO1.
 
 Every proof (except WO6 \<Longrightarrow> WO1 and WO1 \<Longrightarrow> WO2) are described as "clear"
-by Rubin & Rubin (page 2). 
+by Rubin \<and> Rubin (page 2). 
 They refer reader to a book by Gödel to see the proof WO1 \<Longrightarrow> WO2.
 Fortunately order types made this proof also very easy.
 *)
@@ -17,8 +17,8 @@
 (* Auxiliary definitions used in proof *)
 definition
   NN  :: "i => i"  where
-    "NN(y) \<equiv> {m \<in> nat. \<exists>a. \<exists>f. Ord(a) & domain(f)=a  &  
-                        (\<Union>b<a. f`b) = y & (\<forall>b<a. f`b \<lesssim> m)}"
+    "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
@@ -29,9 +29,9 @@
 definition
   vv1 :: "[i, i, i] => i"  where
      "vv1(f,m,b) \<equiv>                                                
-           let g = \<mu> g. (\<exists>d. Ord(d) & (domain(uu(f,b,g,d)) \<noteq> 0 & 
+           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));      
-               d = \<mu> d. domain(uu(f,b,g,d)) \<noteq> 0 &                  
+               d = \<mu> d. domain(uu(f,b,g,d)) \<noteq> 0 \<and>                  
                             domain(uu(f,b,g,d)) \<lesssim> m         
            in  if f`b \<noteq> 0 then domain(uu(f,b,g,d)) else 0"
 
@@ -127,13 +127,13 @@
 (* ********************************************************************** 
     The proof of "WO6 \<Longrightarrow> WO1".  Simplified by L C Paulson.
 
-    From the book "Equivalents of the Axiom of Choice" by Rubin & Rubin,
+    From the book "Equivalents of the Axiom of Choice" by Rubin \<and> Rubin,
     pages 2-5
 ************************************************************************* *)
 
 lemma lt_oadd_odiff_disj:
       "\<lbrakk>k < i++j;  Ord(i);  Ord(j)\<rbrakk> 
-       \<Longrightarrow> k < i  |  (\<not> k<i & k = i ++ (k--i) & (k--i)<j)"
+       \<Longrightarrow> k < i  |  (\<not> k<i \<and> k = i ++ (k--i) \<and> (k--i)<j)"
 apply (rule_tac i = k and j = i in Ord_linear2)
 prefer 4 
    apply (drule odiff_lt_mono2, assumption)
@@ -172,8 +172,8 @@
 lemma cases: 
      "\<forall>b<a. \<forall>g<a. \<forall>d<a. u(f,b,g,d) \<lesssim> m   
       \<Longrightarrow> (\<forall>b<a. f`b \<noteq> 0 \<longrightarrow>  
-                  (\<exists>g<a. \<exists>d<a. u(f,b,g,d) \<noteq> 0 & u(f,b,g,d) \<prec> m))  
-        | (\<exists>b<a. f`b \<noteq> 0 & (\<forall>g<a. \<forall>d<a. u(f,b,g,d) \<noteq> 0 \<longrightarrow>   
+                  (\<exists>g<a. \<exists>d<a. u(f,b,g,d) \<noteq> 0 \<and> u(f,b,g,d) \<prec> m))  
+        | (\<exists>b<a. f`b \<noteq> 0 \<and> (\<forall>g<a. \<forall>d<a. u(f,b,g,d) \<noteq> 0 \<longrightarrow>   
                                         u(f,b,g,d) \<approx> m))"
 apply (unfold lesspoll_def)
 apply (blast del: equalityI)
@@ -210,7 +210,7 @@
 (* ********************************************************************** *)
 lemma nested_LeastI:
      "\<lbrakk>P(a, b);  Ord(a);  Ord(b);   
-         Least_a = (\<mu> a. \<exists>x. Ord(x) & P(a, x))\<rbrakk>   
+         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)
@@ -218,13 +218,13 @@
 done
 
 lemmas nested_Least_instance =
-       nested_LeastI [of "%g d. domain(uu(f,b,g,d)) \<noteq> 0 & 
+       nested_LeastI [of "%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: 
      "\<lbrakk>Ord(a);  m \<in> nat;   
          \<forall>b<a. f`b \<noteq>0 \<longrightarrow>                                        
-             (\<exists>g<a. \<exists>d<a. domain(uu(f,b,g,d)) \<noteq> 0  &                
+             (\<exists>g<a. \<exists>d<a. domain(uu(f,b,g,d)) \<noteq> 0  \<and>                
                           domain(uu(f,b,g,d)) \<lesssim> m);             
          \<forall>b<a. f`b \<lesssim> succ(m);  b<a++a\<rbrakk> 
       \<Longrightarrow> gg1(f,a,m)`b \<lesssim> m"
@@ -439,7 +439,7 @@
 done
 
 (* ********************************************************************** *)
-(* Rubin & Rubin wrote,                                                   *)
+(* Rubin \<and> Rubin wrote,                                                   *)
 (* "It follows from (ii) and mathematical induction that if y*y \<subseteq> y then *)
 (* y can be well-ordered"                                                 *)
 
@@ -471,7 +471,7 @@
 apply (fast elim!: lt_Ord intro: LeastI)
 done
 
-lemma NN_imp_ex_inj: "1 \<in> NN(y) \<Longrightarrow> \<exists>a f. Ord(a) & f \<in> inj(y, a)"
+lemma NN_imp_ex_inj: "1 \<in> NN(y) \<Longrightarrow> \<exists>a f. Ord(a) \<and> f \<in> inj(y, a)"
 apply (unfold NN_def)
 apply (elim CollectE exE conjE)
 apply (rule_tac x = a in exI)
--- a/src/ZF/Arith.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Arith.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -393,7 +393,7 @@
 lemma pred_0 [simp]: "pred(0) = 0"
 by (simp add: pred_def)
 
-lemma eq_succ_imp_eq_m1: "\<lbrakk>i = succ(j); i\<in>nat\<rbrakk> \<Longrightarrow> j = i #- 1 & j \<in>nat"
+lemma eq_succ_imp_eq_m1: "\<lbrakk>i = succ(j); i\<in>nat\<rbrakk> \<Longrightarrow> j = i #- 1 \<and> j \<in>nat"
 by simp
 
 lemma pred_Un_distrib:
@@ -537,7 +537,7 @@
 
 lemma lt_succ_eq_0_disj:
      "\<lbrakk>m\<in>nat; n\<in>nat\<rbrakk>
-      \<Longrightarrow> (m < succ(n)) \<longleftrightarrow> (m = 0 | (\<exists>j\<in>nat. m = succ(j) & j < n))"
+      \<Longrightarrow> (m < succ(n)) \<longleftrightarrow> (m = 0 | (\<exists>j\<in>nat. m = succ(j) \<and> j < n))"
 by (induct_tac "m", auto)
 
 lemma less_diff_conv [rule_format]:
--- a/src/ZF/ArithSimp.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/ArithSimp.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -303,23 +303,23 @@
 apply (simp (no_asm_simp) add: mult_lt_mono2 mult_commute [of _ k])
 done
 
-lemma add_eq_0_iff [iff]: "m#+n = 0 \<longleftrightarrow> natify(m)=0 & natify(n)=0"
-apply (subgoal_tac "natify (m) #+ natify (n) = 0 \<longleftrightarrow> natify (m) =0 & natify (n) =0")
+lemma add_eq_0_iff [iff]: "m#+n = 0 \<longleftrightarrow> natify(m)=0 \<and> natify(n)=0"
+apply (subgoal_tac "natify (m) #+ natify (n) = 0 \<longleftrightarrow> natify (m) =0 \<and> natify (n) =0")
 apply (rule_tac [2] n = "natify (m) " in natE)
  apply (rule_tac [4] n = "natify (n) " in natE)
 apply auto
 done
 
-lemma zero_lt_mult_iff [iff]: "0 < m#*n \<longleftrightarrow> 0 < natify(m) & 0 < natify(n)"
-apply (subgoal_tac "0 < natify (m) #*natify (n) \<longleftrightarrow> 0 < natify (m) & 0 < natify (n) ")
+lemma zero_lt_mult_iff [iff]: "0 < m#*n \<longleftrightarrow> 0 < natify(m) \<and> 0 < natify(n)"
+apply (subgoal_tac "0 < natify (m) #*natify (n) \<longleftrightarrow> 0 < natify (m) \<and> 0 < natify (n) ")
 apply (rule_tac [2] n = "natify (m) " in natE)
  apply (rule_tac [4] n = "natify (n) " in natE)
   apply (rule_tac [3] n = "natify (n) " in natE)
 apply auto
 done
 
-lemma mult_eq_1_iff [iff]: "m#*n = 1 \<longleftrightarrow> natify(m)=1 & natify(n)=1"
-apply (subgoal_tac "natify (m) #* natify (n) = 1 \<longleftrightarrow> natify (m) =1 & natify (n) =1")
+lemma mult_eq_1_iff [iff]: "m#*n = 1 \<longleftrightarrow> natify(m)=1 \<and> natify(n)=1"
+apply (subgoal_tac "natify (m) #* natify (n) = 1 \<longleftrightarrow> natify (m) =1 \<and> natify (n) =1")
 apply (rule_tac [2] n = "natify (m) " in natE)
  apply (rule_tac [4] n = "natify (n) " in natE)
 apply auto
@@ -342,7 +342,7 @@
 subsection\<open>Cancellation Laws for Common Factors in Comparisons\<close>
 
 lemma mult_less_cancel_lemma:
-     "\<lbrakk>k: nat; m: nat; n: nat\<rbrakk> \<Longrightarrow> (m#*k < n#*k) \<longleftrightarrow> (0<k & m<n)"
+     "\<lbrakk>k: nat; m: nat; n: nat\<rbrakk> \<Longrightarrow> (m#*k < n#*k) \<longleftrightarrow> (0<k \<and> m<n)"
 apply (safe intro!: mult_lt_mono1)
 apply (erule natE, auto)
 apply (rule not_le_iff_lt [THEN iffD1])
@@ -351,13 +351,13 @@
 done
 
 lemma mult_less_cancel2 [simp]:
-     "(m#*k < n#*k) \<longleftrightarrow> (0 < natify(k) & natify(m) < natify(n))"
+     "(m#*k < n#*k) \<longleftrightarrow> (0 < natify(k) \<and> natify(m) < natify(n))"
 apply (rule iff_trans)
 apply (rule_tac [2] mult_less_cancel_lemma, auto)
 done
 
 lemma mult_less_cancel1 [simp]:
-     "(k#*m < k#*n) \<longleftrightarrow> (0 < natify(k) & natify(m) < natify(n))"
+     "(k#*m < k#*n) \<longleftrightarrow> (0 < natify(k) \<and> natify(m) < natify(n))"
 apply (simp (no_asm) add: mult_less_cancel2 mult_commute [of k])
 done
 
@@ -374,7 +374,7 @@
 lemma mult_le_cancel_le1: "k \<in> nat \<Longrightarrow> k #* m \<le> k \<longleftrightarrow> (0 < k \<longrightarrow> natify(m) \<le> 1)"
 by (cut_tac k = k and m = m and n = 1 in mult_le_cancel1, auto)
 
-lemma Ord_eq_iff_le: "\<lbrakk>Ord(m); Ord(n)\<rbrakk> \<Longrightarrow> m=n \<longleftrightarrow> (m \<le> n & n \<le> m)"
+lemma Ord_eq_iff_le: "\<lbrakk>Ord(m); Ord(n)\<rbrakk> \<Longrightarrow> m=n \<longleftrightarrow> (m \<le> n \<and> n \<le> m)"
 by (blast intro: le_anti_sym)
 
 lemma mult_cancel2_lemma:
@@ -523,7 +523,7 @@
 
 lemma raw_nat_diff_split:
      "\<lbrakk>a:nat; b:nat\<rbrakk> \<Longrightarrow>
-      (P(a #- b)) \<longleftrightarrow> ((a < b \<longrightarrow>P(0)) & (\<forall>d\<in>nat. a = b #+ d \<longrightarrow> P(d)))"
+      (P(a #- b)) \<longleftrightarrow> ((a < b \<longrightarrow>P(0)) \<and> (\<forall>d\<in>nat. a = b #+ d \<longrightarrow> P(d)))"
 apply (case_tac "a < b")
  apply (force simp add: nat_lt_imp_diff_eq_0)
 apply (rule iffI, force, simp)
@@ -533,7 +533,7 @@
 
 lemma nat_diff_split:
    "(P(a #- b)) \<longleftrightarrow>
-    (natify(a) < natify(b) \<longrightarrow>P(0)) & (\<forall>d\<in>nat. natify(a) = b #+ d \<longrightarrow> P(d))"
+    (natify(a) < natify(b) \<longrightarrow>P(0)) \<and> (\<forall>d\<in>nat. natify(a) = b #+ d \<longrightarrow> P(d))"
 apply (cut_tac P=P and a="natify(a)" and b="natify(b)" in raw_nat_diff_split)
 apply simp_all
 done
--- a/src/ZF/Bin.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Bin.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -377,7 +377,7 @@
 
 lemma iszero_integ_of_BIT:
      "\<lbrakk>w \<in> bin; x \<in> bool\<rbrakk>
-      \<Longrightarrow> iszero (integ_of (w BIT x)) \<longleftrightarrow> (x=0 & iszero (integ_of(w)))"
+      \<Longrightarrow> iszero (integ_of (w BIT x)) \<longleftrightarrow> (x=0 \<and> iszero (integ_of(w)))"
 apply (unfold iszero_def, simp)
 apply (subgoal_tac "integ_of (w) \<in> int")
 apply typecheck
@@ -578,7 +578,7 @@
 apply (blast intro: zless_zle_trans)
 done
 
-lemma zless_nat_conj: "(nat_of(w) < nat_of(z)) \<longleftrightarrow> ($#0 $< z & w $< z)"
+lemma zless_nat_conj: "(nat_of(w) < nat_of(z)) \<longleftrightarrow> ($#0 $< z \<and> w $< z)"
 apply (case_tac "$#0 $< z")
 apply (auto simp add: zless_nat_conj_lemma nat_le_int0 not_zless_iff_zle)
 done
--- a/src/ZF/Cardinal.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Cardinal.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -10,7 +10,7 @@
 definition
   (*least ordinal operator*)
    Least    :: "(i=>o) => i"    (binder \<open>\<mu> \<close> 10)  where
-     "Least(P) \<equiv> THE i. Ord(i) & P(i) & (\<forall>j. j<i \<longrightarrow> \<not>P(j))"
+     "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
@@ -22,7 +22,7 @@
 
 definition
   lesspoll :: "[i,i] => o"     (infixl \<open>\<prec>\<close> 50)  where
-    "A \<prec> B \<equiv> A \<lesssim> B & \<not>(A \<approx> B)"
+    "A \<prec> B \<equiv> A \<lesssim> B \<and> \<not>(A \<approx> B)"
 
 definition
   cardinal :: "i=>i"           (\<open>|_|\<close>)  where
@@ -56,9 +56,9 @@
 
 lemma decomposition:
      "\<lbrakk>f \<in> X->Y;  g \<in> Y->X\<rbrakk> \<Longrightarrow>
-      \<exists>XA XB YA YB. (XA \<inter> XB = 0) & (XA \<union> XB = X) &
-                      (YA \<inter> YB = 0) & (YA \<union> YB = Y) &
-                      f``XA=YA & g``YB=XB"
+      \<exists>XA XB YA YB. (XA \<inter> XB = 0) \<and> (XA \<union> XB = X) \<and>
+                      (YA \<inter> YB = 0) \<and> (YA \<union> YB = Y) \<and>
+                      f``XA=YA \<and> g``YB=XB"
 apply (intro exI conjI)
 apply (rule_tac [6] Banach_last_equation)
 apply (rule_tac [5] refl)
@@ -134,7 +134,7 @@
     "\<lbrakk>X \<approx> Y; \<lbrakk>X \<lesssim> Y; Y \<lesssim> X\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
 by (blast intro: eqpoll_imp_lepoll eqpoll_sym)
 
-lemma eqpoll_iff: "X \<approx> Y \<longleftrightarrow> X \<lesssim> Y & Y \<lesssim> X"
+lemma eqpoll_iff: "X \<approx> Y \<longleftrightarrow> X \<lesssim> Y \<and> Y \<lesssim> X"
 by (blast intro: eqpollI elim!: eqpollE)
 
 lemma lepoll_0_is_0: "A \<lesssim> 0 \<Longrightarrow> A = 0"
@@ -303,13 +303,13 @@
 
 (*If there is no such P then \<mu> is vacuously 0*)
 lemma Least_0:
-    "\<lbrakk>\<not> (\<exists>i. Ord(i) & P(i))\<rbrakk> \<Longrightarrow> (\<mu> x. P(x)) = 0"
+    "\<lbrakk>\<not> (\<exists>i. Ord(i) \<and> P(i))\<rbrakk> \<Longrightarrow> (\<mu> x. P(x)) = 0"
 apply (unfold Least_def)
 apply (rule the_0, blast)
 done
 
 lemma Ord_Least [intro,simp,TC]: "Ord(\<mu> x. P(x))"
-proof (cases "\<exists>i. Ord(i) & P(i)")
+proof (cases "\<exists>i. Ord(i) \<and> P(i)")
   case True 
   then obtain i where "P(i)" "Ord(i)"  by auto
   hence " (\<mu> x. P(x)) \<le> i"  by (rule Least_le) 
@@ -401,7 +401,7 @@
 done
 
 text\<open>The cardinals are the initial ordinals.\<close>
-lemma Card_iff_initial: "Card(K) \<longleftrightarrow> Ord(K) & (\<forall>j. j<K \<longrightarrow> \<not> j \<approx> K)"
+lemma Card_iff_initial: "Card(K) \<longleftrightarrow> Ord(K) \<and> (\<forall>j. j<K \<longrightarrow> \<not> j \<approx> K)"
 proof -
   { fix j
     assume K: "Card(K)" "j \<approx> K"
@@ -438,7 +438,7 @@
 lemma Card_cardinal [iff]: "Card(|A|)"
 proof (unfold cardinal_def)
   show "Card(\<mu> i. i \<approx> A)"
-    proof (cases "\<exists>i. Ord (i) & i \<approx> A")
+    proof (cases "\<exists>i. Ord (i) \<and> i \<approx> A")
       case False thus ?thesis           \<comment> \<open>degenerate case\<close>
         by (simp add: Least_0 Card_0)
     next
@@ -620,7 +620,7 @@
 by (rule nat_lepoll_imp_le [THEN lt_irrefl], auto)
 
 lemma nat_lepoll_imp_ex_eqpoll_n:
-     "\<lbrakk>n \<in> nat;  nat \<lesssim> X\<rbrakk> \<Longrightarrow> \<exists>Y. Y \<subseteq> X & n \<approx> Y"
+     "\<lbrakk>n \<in> nat;  nat \<lesssim> X\<rbrakk> \<Longrightarrow> \<exists>Y. Y \<subseteq> X \<and> n \<approx> Y"
 apply (unfold lepoll_def eqpoll_def)
 apply (fast del: subsetI subsetCE
             intro!: subset_SIs
@@ -962,7 +962,7 @@
 by (fast elim!: eqpoll_sym [THEN eqpoll_0_is_0, THEN succ_neq_0])
 
 lemma Finite_Fin_lemma [rule_format]:
-     "n \<in> nat \<Longrightarrow> \<forall>A. (A\<approx>n & A \<subseteq> X) \<longrightarrow> A \<in> Fin(X)"
+     "n \<in> nat \<Longrightarrow> \<forall>A. (A\<approx>n \<and> A \<subseteq> X) \<longrightarrow> A \<in> Fin(X)"
 apply (induct_tac n)
 apply (rule allI)
 apply (fast intro!: Fin.emptyI dest!: eqpoll_imp_lepoll [THEN lepoll_0_is_0])
@@ -1014,7 +1014,7 @@
           intro: Un_upper1 [THEN Fin_mono, THEN subsetD]
                  Un_upper2 [THEN Fin_mono, THEN subsetD])
 
-lemma Finite_Un_iff [simp]: "Finite(A \<union> B) \<longleftrightarrow> (Finite(A) & Finite(B))"
+lemma Finite_Un_iff [simp]: "Finite(A \<union> B) \<longleftrightarrow> (Finite(A) \<and> Finite(B))"
 by (blast intro: subset_Finite Finite_Un)
 
 text\<open>The converse must hold too.\<close>
@@ -1040,7 +1040,7 @@
 apply (case_tac "a \<in> A")
 apply (subgoal_tac [2] "A-{a}=A", auto)
 apply (rule_tac x = "succ (n) " in bexI)
-apply (subgoal_tac "cons (a, A - {a}) = A & cons (n, n) = succ (n) ")
+apply (subgoal_tac "cons (a, A - {a}) = A \<and> cons (n, n) = succ (n) ")
 apply (drule_tac a = a and b = n in cons_eqpoll_cong)
 apply (auto dest: mem_irrefl)
 done
--- a/src/ZF/CardinalArith.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/CardinalArith.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -9,7 +9,7 @@
 
 definition
   InfCard       :: "i=>o"  where
-    "InfCard(i) \<equiv> Card(i) & nat \<le> i"
+    "InfCard(i) \<equiv> Card(i) \<and> nat \<le> i"
 
 definition
   cmult         :: "[i,i]=>i"       (infixl \<open>\<otimes>\<close> 70)  where
@@ -31,13 +31,13 @@
     \<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) & z = ordertype(X,r)}"
+         \<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
     \<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) & K<L"
+    "csucc(K) \<equiv> \<mu> L. Card(L) \<and> K<L"
 
 
 lemma Card_Union [simp,intro,TC]:
@@ -48,7 +48,7 @@
 next
   fix j
   assume j: "j < \<Union>A"
-  hence "\<exists>c\<in>A. j < c & Card(c)" using A
+  hence "\<exists>c\<in>A. j < c \<and> Card(c)" using A
     by (auto simp add: lt_def intro: Card_is_Ord)
   then obtain c where c: "c\<in>A" "j < c" "Card(c)"
     by blast
@@ -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 & y \<le> z"
+ "\<lbrakk><<x,y>, <z,z>> \<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)
@@ -510,17 +510,17 @@
 lemma csquare_ltI:
  "\<lbrakk>x<z;  y<z;  z<K\<rbrakk> \<Longrightarrow>  <<x,y>, <z,z>> \<in> csquare_rel(K)"
 apply (unfold csquare_rel_def)
-apply (subgoal_tac "x<K & y<K")
+apply (subgoal_tac "x<K \<and> y<K")
  prefer 2 apply (blast intro: lt_trans)
 apply (elim ltE)
 apply (simp add: rvimage_iff Un_absorb Un_least_mem_iff ltD)
 done
 
-(*Part of the traditional proof.  UNUSED since it's harder to prove & apply *)
+(*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 & y=z"
+ "\<lbrakk>x \<le> z;  y \<le> z;  z<K\<rbrakk> \<Longrightarrow> <<x,y>, <z,z>> \<in> csquare_rel(K) | x=z \<and> y=z"
 apply (unfold csquare_rel_def)
-apply (subgoal_tac "x<K & y<K")
+apply (subgoal_tac "x<K \<and> y<K")
  prefer 2 apply (blast intro: lt_trans1)
 apply (elim ltE)
 apply (simp add: rvimage_iff Un_absorb Un_least_mem_iff ltD)
@@ -535,7 +535,7 @@
       "\<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>"
-apply (subgoal_tac "z<K & well_ord (K*K, csquare_rel (K))")
+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)
 apply (rule csquare_ltI [THEN ordermap_mono, THEN ltI])
@@ -745,7 +745,7 @@
 (*Allows selective unfolding.  Less work than deriving intro/elim rules*)
 lemma jump_cardinal_iff:
      "i \<in> jump_cardinal(K) \<longleftrightarrow>
-      (\<exists>r X. r \<subseteq> K*K & X \<subseteq> K & well_ord(X,r) & i = ordertype(X,r))"
+      (\<exists>r X. r \<subseteq> K*K \<and> X \<subseteq> K \<and> well_ord(X,r) \<and> i = ordertype(X,r))"
 apply (unfold jump_cardinal_def)
 apply (blast del: subsetI)
 done
@@ -787,7 +787,7 @@
 
 subsection\<open>Basic Properties of Successor Cardinals\<close>
 
-lemma csucc_basic: "Ord(K) \<Longrightarrow> Card(csucc(K)) & K < csucc(K)"
+lemma csucc_basic: "Ord(K) \<Longrightarrow> Card(csucc(K)) \<and> K < csucc(K)"
 apply (unfold csucc_def)
 apply (rule LeastI)
 apply (blast intro: Card_jump_cardinal K_lt_jump_cardinal Ord_jump_cardinal)+
--- a/src/ZF/Coind/ECR.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Coind/ECR.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -28,7 +28,7 @@
 definition
   hastyenv :: "[i,i] => o"  where
     "hastyenv(ve,te) \<equiv>                        
-         ve_dom(ve) = te_dom(te) &          
+         ve_dom(ve) = te_dom(te) \<and>          
          (\<forall>x \<in> ve_dom(ve). <ve_app(ve,x),te_app(te,x)> \<in> HasTyRel)"
 
 (* Specialised co-induction rule *)
--- a/src/ZF/Coind/Static.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Coind/Static.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -17,9 +17,9 @@
 definition
   isofenv :: "[i,i] => o"  where
    "isofenv(ve,te) \<equiv>                
-      ve_dom(ve) = te_dom(te) &            
+      ve_dom(ve) = te_dom(te) \<and>            
       (\<forall>x \<in> ve_dom(ve).                          
-        \<exists>c \<in> Const. ve_app(ve,x) = v_const(c) & isof(c,te_app(te,x)))"
+        \<exists>c \<in> Const. ve_app(ve,x) = v_const(c) \<and> isof(c,te_app(te,x)))"
   
 
 (*** Elaboration ***)
--- a/src/ZF/Constructible/AC_in_L.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/AC_in_L.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -56,7 +56,7 @@
       { fix y ys
         assume "y \<in> A" and "ys \<in> list(A)"
         with Cons
-        have "\<langle>Cons(x,xs),Cons(y,ys)\<rangle> \<in> rlist(A,r) \<or> x=y & xs = ys \<or> \<langle>Cons(y,ys), Cons(x,xs)\<rangle> \<in> rlist(A,r)" 
+        have "\<langle>Cons(x,xs),Cons(y,ys)\<rangle> \<in> rlist(A,r) \<or> x=y \<and> xs = ys \<or> \<langle>Cons(y,ys), Cons(x,xs)\<rangle> \<in> rlist(A,r)" 
           apply (rule_tac i = "length(xs)" and j = "length(ys)" in Ord_linear_lt)
           apply (simp_all add: shorterI)
           apply (rule linearE [OF r, of x y]) 
@@ -172,7 +172,7 @@
 
 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 & y=v)"
+     \<Longrightarrow> (fn`<x,y> = fn`<u,v>) \<longleftrightarrow> (x=u \<and> y=v)"
 by (blast dest: inj_apply_equality [OF fn_inj])
 
 lemma (in Nat_Times_Nat) enum_type [TC,simp]:
@@ -247,8 +247,8 @@
     \<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)) &
-             X = {x\<in>A. sats(A, p, Cons(x,env))} &
+             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"
 
 definition
@@ -335,9 +335,9 @@
     "rlimit(i,r) \<equiv>
        if Limit(i) then 
          {z: Lset(i) * Lset(i).
-          \<exists>x' x. z = <x',x> &
+          \<exists>x' x. z = <x',x> \<and>
                  (lrank(x') < lrank(x) |
-                  (lrank(x') = lrank(x) & <x',x> \<in> r(succ(lrank(x)))))}
+                  (lrank(x') = lrank(x) \<and> <x',x> \<in> r(succ(lrank(x)))))}
        else 0"
 
 definition
--- a/src/ZF/Constructible/DPow_absolute.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/DPow_absolute.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -17,9 +17,9 @@
 
 (* is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o"
    "is_formula_rec(M,MH,p,z)  \<equiv>
-      \<exists>dp[M]. \<exists>i[M]. \<exists>f[M]. finite_ordinal(M,dp) & is_depth(M,p,dp) & 
+      \<exists>dp[M]. \<exists>i[M]. \<exists>f[M]. finite_ordinal(M,dp) \<and> is_depth(M,p,dp) \<and> 
        2       1      0
-             successor(M,dp,i) & fun_apply(M,f,p,z) & is_transrec(M,MH,i,f)"
+             successor(M,dp,i) \<and> fun_apply(M,f,p,z) \<and> is_transrec(M,MH,i,f)"
 *)
 
 definition
@@ -113,7 +113,7 @@
 lemma DPow'_eq: 
   "DPow'(A) = {z . ep \<in> list(A) * formula, 
                     \<exists>env \<in> list(A). \<exists>p \<in> formula. 
-                       ep = <env,p> & z = {x\<in>A. sats(A, p, Cons(x,env))}}"
+                       ep = <env,p> \<and> z = {x\<in>A. sats(A, p, Cons(x,env))}}"
 by (simp add: DPow'_def, blast) 
 
 
@@ -192,8 +192,8 @@
  and rep: 
     "M(A)
     \<Longrightarrow> strong_replacement (M, 
-         \<lambda>ep z. \<exists>env[M]. \<exists>p[M]. mem_formula(M,p) & mem_list(M,A,env) &
-                  pair(M,env,p,ep) & 
+         \<lambda>ep z. \<exists>env[M]. \<exists>p[M]. mem_formula(M,p) \<and> mem_list(M,A,env) \<and>
+                  pair(M,env,p,ep) \<and> 
                   is_Collect(M, A, \<lambda>x. is_DPow_sats(M,A,env,p,x), z))"
 
 lemma (in M_DPow) sep':
@@ -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> & z = {x \<in> A . sats(A, p, Cons(x, env))})" 
+                  ep = <env,p> \<and> z = {x \<in> A . sats(A, p, Cons(x, env))})" 
 by (insert rep [of A], simp add: Collect_DPow_sats_abs) 
 
 
@@ -223,8 +223,8 @@
   is_DPow' :: "[i=>o,i,i] => o" where
     "is_DPow'(M,A,Z) \<equiv> 
        \<forall>X[M]. X \<in> Z \<longleftrightarrow> 
-         subset(M,X,A) & 
-           (\<exists>env[M]. \<exists>p[M]. mem_formula(M,p) & mem_list(M,A,env) &
+         subset(M,X,A) \<and> 
+           (\<exists>env[M]. \<exists>p[M]. mem_formula(M,p) \<and> mem_list(M,A,env) \<and>
                     is_Collect(M, A, is_DPow_sats(M,A,env,p), X))"
 
 lemma (in M_DPow) DPow'_abs:
@@ -255,14 +255,14 @@
 subsubsection\<open>The Instance of Replacement\<close>
 
 lemma DPow_replacement_Reflects:
- "REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B &
+ "REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B \<and>
              (\<exists>env[L]. \<exists>p[L].
-               mem_formula(L,p) & mem_list(L,A,env) & pair(L,env,p,u) &
+               mem_formula(L,p) \<and> mem_list(L,A,env) \<and> pair(L,env,p,u) \<and>
                is_Collect (L, A, is_DPow_sats(L,A,env,p), x)),
-    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B &
+    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and>
              (\<exists>env \<in> Lset(i). \<exists>p \<in> Lset(i).
-               mem_formula(##Lset(i),p) & mem_list(##Lset(i),A,env) & 
-               pair(##Lset(i),env,p,u) &
+               mem_formula(##Lset(i),p) \<and> mem_list(##Lset(i),A,env) \<and> 
+               pair(##Lset(i),env,p,u) \<and>
                is_Collect (##Lset(i), A, is_DPow_sats(##Lset(i),A,env,p), x))]"
 apply (unfold is_Collect_def) 
 apply (intro FOL_reflections function_reflections mem_formula_reflection
@@ -272,8 +272,8 @@
 lemma DPow_replacement:
     "L(A)
     \<Longrightarrow> strong_replacement (L, 
-         \<lambda>ep z. \<exists>env[L]. \<exists>p[L]. mem_formula(L,p) & mem_list(L,A,env) &
-                  pair(L,env,p,ep) & 
+         \<lambda>ep z. \<exists>env[L]. \<exists>p[L]. mem_formula(L,p) \<and> mem_list(L,A,env) \<and>
+                  pair(L,env,p,ep) \<and> 
                   is_Collect(L, A, \<lambda>x. is_DPow_sats(L,A,env,p,x), z))"
 apply (rule strong_replacementI)
 apply (rule_tac u="{A,B}" 
@@ -309,7 +309,7 @@
 enclosed within a single quantifier.\<close>
 
 (* is_Collect :: "[i=>o,i,i=>o,i] => o"
-    "is_Collect(M,A,P,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> A & P(x)" *)
+    "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
@@ -360,7 +360,7 @@
  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(M,A,P,z) \<equiv> \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,u))" *)
+    "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
@@ -412,8 +412,8 @@
 
 (*  "is_DPow'(M,A,Z) \<equiv> 
        \<forall>X[M]. X \<in> Z \<longleftrightarrow> 
-         subset(M,X,A) & 
-           (\<exists>env[M]. \<exists>p[M]. mem_formula(M,p) & mem_list(M,A,env) &
+         subset(M,X,A) \<and> 
+           (\<exists>env[M]. \<exists>p[M]. mem_formula(M,p) \<and> mem_list(M,A,env) \<and>
                     is_Collect(M, A, is_DPow_sats(M,A,env,p), X))" *)
 
 definition
@@ -458,11 +458,11 @@
 definition
   transrec_body :: "[i=>o,i,i,i,i] => o" where
     "transrec_body(M,g,x) \<equiv>
-      \<lambda>y z. \<exists>gy[M]. y \<in> x & fun_apply(M,g,y,gy) & is_DPow'(M,gy,z)"
+      \<lambda>y z. \<exists>gy[M]. y \<in> x \<and> fun_apply(M,g,y,gy) \<and> is_DPow'(M,gy,z)"
 
 lemma (in M_DPow) transrec_body_abs:
      "\<lbrakk>M(x); M(g); M(z)\<rbrakk>
-    \<Longrightarrow> transrec_body(M,g,x,y,z) \<longleftrightarrow> y \<in> x & z = DPow'(g`y)"
+    \<Longrightarrow> transrec_body(M,g,x,y,z) \<longleftrightarrow> y \<in> x \<and> z = DPow'(g`y)"
 by (simp add: transrec_body_def DPow'_abs transM [of _ x])
 
 locale M_Lset = M_DPow +
@@ -470,13 +470,13 @@
    "\<lbrakk>M(x); M(g)\<rbrakk> \<Longrightarrow> strong_replacement(M, \<lambda>y z. transrec_body(M,g,x,y,z))"
  and transrec_rep: 
     "M(i) \<Longrightarrow> transrec_replacement(M, \<lambda>x f u. 
-              \<exists>r[M]. is_Replace(M, x, transrec_body(M,f,x), r) & 
+              \<exists>r[M]. is_Replace(M, x, transrec_body(M,f,x), r) \<and> 
                      big_union(M, r, u), i)"
 
 
 lemma (in M_Lset) strong_rep':
    "\<lbrakk>M(x); M(g)\<rbrakk>
-    \<Longrightarrow> strong_replacement(M, \<lambda>y z. y \<in> x & z = DPow'(g`y))"
+    \<Longrightarrow> strong_replacement(M, \<lambda>y z. y \<in> x \<and> z = DPow'(g`y))"
 by (insert strong_rep [of x g], simp add: transrec_body_abs)
 
 lemma (in M_Lset) DPow_apply_closed:
@@ -535,10 +535,10 @@
 subsubsection\<open>The First Instance of Replacement\<close>
 
 lemma strong_rep_Reflects:
- "REFLECTS [\<lambda>u. \<exists>v[L]. v \<in> B & (\<exists>gy[L].
-                          v \<in> x & fun_apply(L,g,v,gy) & is_DPow'(L,gy,u)),
-   \<lambda>i u. \<exists>v \<in> Lset(i). v \<in> B & (\<exists>gy \<in> Lset(i).
-            v \<in> x & fun_apply(##Lset(i),g,v,gy) & is_DPow'(##Lset(i),gy,u))]"
+ "REFLECTS [\<lambda>u. \<exists>v[L]. v \<in> B \<and> (\<exists>gy[L].
+                          v \<in> x \<and> fun_apply(L,g,v,gy) \<and> is_DPow'(L,gy,u)),
+   \<lambda>i u. \<exists>v \<in> Lset(i). v \<in> B \<and> (\<exists>gy \<in> Lset(i).
+            v \<in> x \<and> fun_apply(##Lset(i),g,v,gy) \<and> is_DPow'(##Lset(i),gy,u))]"
 by (intro FOL_reflections function_reflections DPow'_reflection)
 
 lemma strong_rep:
@@ -555,18 +555,18 @@
 subsubsection\<open>The Second Instance of Replacement\<close>
 
 lemma transrec_rep_Reflects:
- "REFLECTS [\<lambda>x. \<exists>v[L]. v \<in> B &
-              (\<exists>y[L]. pair(L,v,y,x) &
+ "REFLECTS [\<lambda>x. \<exists>v[L]. v \<in> B \<and>
+              (\<exists>y[L]. pair(L,v,y,x) \<and>
              is_wfrec (L, \<lambda>x f u. \<exists>r[L].
                 is_Replace (L, x, \<lambda>y z. 
-                     \<exists>gy[L]. y \<in> x & fun_apply(L,f,y,gy) & 
-                      is_DPow'(L,gy,z), r) & big_union(L,r,u), mr, v, y)),
-    \<lambda>i x. \<exists>v \<in> Lset(i). v \<in> B &
-              (\<exists>y \<in> Lset(i). pair(##Lset(i),v,y,x) &
+                     \<exists>gy[L]. y \<in> x \<and> fun_apply(L,f,y,gy) \<and> 
+                      is_DPow'(L,gy,z), r) \<and> big_union(L,r,u), mr, v, y)),
+    \<lambda>i x. \<exists>v \<in> Lset(i). v \<in> B \<and>
+              (\<exists>y \<in> Lset(i). pair(##Lset(i),v,y,x) \<and>
              is_wfrec (##Lset(i), \<lambda>x f u. \<exists>r \<in> Lset(i).
                 is_Replace (##Lset(i), x, \<lambda>y z. 
-                     \<exists>gy \<in> Lset(i). y \<in> x & fun_apply(##Lset(i),f,y,gy) & 
-                      is_DPow'(##Lset(i),gy,z), r) & 
+                     \<exists>gy \<in> Lset(i). y \<in> x \<and> fun_apply(##Lset(i),f,y,gy) \<and> 
+                      is_DPow'(##Lset(i),gy,z), r) \<and> 
                       big_union(##Lset(i),r,u), mr, v, y))]" 
 apply (simp only: rex_setclass_is_bex [symmetric])
   \<comment> \<open>Convert \<open>\<exists>y\<in>Lset(i)\<close> to \<open>\<exists>y[##Lset(i)]\<close> within the body
@@ -579,7 +579,7 @@
 lemma transrec_rep: 
     "\<lbrakk>L(j)\<rbrakk>
     \<Longrightarrow> transrec_replacement(L, \<lambda>x f u. 
-              \<exists>r[L]. is_Replace(L, x, transrec_body(L,f,x), r) & 
+              \<exists>r[L]. is_Replace(L, x, transrec_body(L,f,x), r) \<and> 
                      big_union(L, r, u), j)"
 apply (rule L.transrec_replacementI, assumption)
 apply (unfold transrec_body_def)  
@@ -615,7 +615,7 @@
 definition
   constructible :: "[i=>o,i] => o" where
     "constructible(M,x) \<equiv>
-       \<exists>i[M]. \<exists>Li[M]. ordinal(M,i) & is_Lset(M,i,Li) & x \<in> Li"
+       \<exists>i[M]. \<exists>Li[M]. ordinal(M,i) \<and> is_Lset(M,i,Li) \<and> x \<in> Li"
 
 theorem V_equals_L_in_L:
   "L(x) \<longleftrightarrow> constructible(L,x)"
--- a/src/ZF/Constructible/Datatype_absolute.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/Datatype_absolute.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -11,7 +11,7 @@
 
 definition
   directed :: "i=>o" where
-   "directed(A) \<equiv> A\<noteq>0 & (\<forall>x\<in>A. \<forall>y\<in>A. x \<union> y \<in> A)"
+   "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
@@ -116,13 +116,13 @@
 definition
   iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => 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) & isF(gm,u),
+        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(M,isF,v,n,Z) \<equiv> 
-      \<exists>sn[M]. \<exists>msn[M]. successor(M,n,sn) & membership(M,sn,msn) &
+      \<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
@@ -212,7 +212,7 @@
   is_list_functor :: "[i=>o,i,i,i] => o" where
     "is_list_functor(M,A,X,Z) \<equiv> 
         \<exists>n1[M]. \<exists>AX[M]. 
-         number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)"
+         number1(M,n1) \<and> cartprod(M,A,X,AX) \<and> is_sum(M,n1,AX,Z)"
 
 lemma (in M_basic) list_functor_abs [simp]: 
      "\<lbrakk>M(A); M(X); M(Z)\<rbrakk> \<Longrightarrow> is_list_functor(M,A,X,Z) \<longleftrightarrow> (Z = {0} + A*X)"
@@ -265,9 +265,9 @@
   is_formula_functor :: "[i=>o,i,i] => 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') & cartprod(M,nat',nat',natnat) & 
-          is_sum(M,natnat,natnat,natnatsum) &
-          cartprod(M,X,X,XX) & is_sum(M,XX,X,X3) & 
+          omega(M,nat') \<and> cartprod(M,nat',nat',natnat) \<and> 
+          is_sum(M,natnat,natnat,natnatsum) \<and>
+          cartprod(M,X,X,XX) \<and> is_sum(M,XX,X,X3) \<and> 
           is_sum(M,natnatsum,X3,Z)"
 
 lemma (in M_trancl) formula_functor_abs [simp]: 
@@ -287,7 +287,7 @@
 by (simp add: list_N_def Nil_def)
 
 lemma Cons_in_list_N [simp]:
-     "Cons(a,l) \<in> list_N(A,succ(n)) \<longleftrightarrow> a\<in>A & l \<in> list_N(A,n)"
+     "Cons(a,l) \<in> list_N(A,succ(n)) \<longleftrightarrow> a\<in>A \<and> l \<in> list_N(A,n)"
 by (simp add: list_N_def Cons_def) 
 
 text\<open>These two aren't simprules because they reveal the underlying
@@ -343,14 +343,14 @@
 definition
   is_list_N :: "[i=>o,i,i,i] => o" where
     "is_list_N(M,A,n,Z) \<equiv> 
-      \<exists>zero[M]. empty(M,zero) & 
+      \<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(M,A,l) \<equiv> 
       \<exists>n[M]. \<exists>listn[M]. 
-       finite_ordinal(M,n) & is_list_N(M,A,n,listn) & l \<in> listn"
+       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
@@ -374,15 +374,15 @@
     "formula_N(n) \<equiv> (\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X)) ^ n (0)"
 
 lemma Member_in_formula_N [simp]:
-     "Member(x,y) \<in> formula_N(succ(n)) \<longleftrightarrow> x \<in> nat & y \<in> nat"
+     "Member(x,y) \<in> formula_N(succ(n)) \<longleftrightarrow> x \<in> nat \<and> y \<in> nat"
 by (simp add: formula_N_def Member_def) 
 
 lemma Equal_in_formula_N [simp]:
-     "Equal(x,y) \<in> formula_N(succ(n)) \<longleftrightarrow> x \<in> nat & y \<in> nat"
+     "Equal(x,y) \<in> formula_N(succ(n)) \<longleftrightarrow> x \<in> nat \<and> y \<in> nat"
 by (simp add: formula_N_def Equal_def) 
 
 lemma Nand_in_formula_N [simp]:
-     "Nand(x,y) \<in> formula_N(succ(n)) \<longleftrightarrow> x \<in> formula_N(n) & y \<in> formula_N(n)"
+     "Nand(x,y) \<in> formula_N(succ(n)) \<longleftrightarrow> x \<in> formula_N(n) \<and> y \<in> formula_N(n)"
 by (simp add: formula_N_def Nand_def) 
 
 lemma Forall_in_formula_N [simp]:
@@ -447,7 +447,7 @@
 definition
   is_formula_N :: "[i=>o,i,i] => o" where
     "is_formula_N(M,n,Z) \<equiv> 
-      \<exists>zero[M]. empty(M,zero) & 
+      \<exists>zero[M]. empty(M,zero) \<and> 
                 is_iterates(M, is_formula_functor(M), zero, n, Z)"
 
 
@@ -455,7 +455,7 @@
   mem_formula :: "[i=>o,i] => o" where
     "mem_formula(M,p) \<equiv> 
       \<exists>n[M]. \<exists>formn[M]. 
-       finite_ordinal(M,n) & is_formula_N(M,n,formn) & p \<in> formn"
+       finite_ordinal(M,n) \<and> is_formula_N(M,n,formn) \<and> p \<in> formn"
 
 definition
   is_formula :: "[i=>o,i] => o" where
@@ -466,12 +466,12 @@
    "M(A) \<Longrightarrow> iterates_replacement(M, is_list_functor(M,A), 0)"
   and list_replacement2:
    "M(A) \<Longrightarrow> strong_replacement(M,
-         \<lambda>n y. n\<in>nat & is_iterates(M, is_list_functor(M,A), 0, n, y))"
+         \<lambda>n y. n\<in>nat \<and> is_iterates(M, is_list_functor(M,A), 0, n, y))"
   and formula_replacement1:
    "iterates_replacement(M, is_formula_functor(M), 0)"
   and formula_replacement2:
    "strong_replacement(M,
-         \<lambda>n y. n\<in>nat & is_iterates(M, is_formula_functor(M), 0, n, y))"
+         \<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)"
 
@@ -479,7 +479,7 @@
 subsubsection\<open>Absoluteness of the List Construction\<close>
 
 lemma (in M_datatypes) list_replacement2':
-  "M(A) \<Longrightarrow> strong_replacement(M, \<lambda>n y. n\<in>nat & y = (\<lambda>X. {0} + A * X)^n (0))"
+  "M(A) \<Longrightarrow> strong_replacement(M, \<lambda>n y. n\<in>nat \<and> y = (\<lambda>X. {0} + A * X)^n (0))"
 apply (insert list_replacement2 [of A])
 apply (rule strong_replacement_cong [THEN iffD1])
 apply (rule conj_cong [OF iff_refl iterates_abs [of "is_list_functor(M,A)"]])
@@ -527,7 +527,7 @@
 subsubsection\<open>Absoluteness of Formulas\<close>
 
 lemma (in M_datatypes) formula_replacement2':
-  "strong_replacement(M, \<lambda>n y. n\<in>nat & y = (\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X))^n (0))"
+  "strong_replacement(M, \<lambda>n y. n\<in>nat \<and> y = (\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X))^n (0))"
 apply (insert formula_replacement2)
 apply (rule strong_replacement_cong [THEN iffD1])
 apply (rule conj_cong [OF iff_refl iterates_abs [of "is_formula_functor(M)"]])
@@ -595,7 +595,7 @@
   mem_eclose :: "[i=>o,i,i] => o" where
     "mem_eclose(M,A,l) \<equiv>
       \<exists>n[M]. \<exists>eclosen[M].
-       finite_ordinal(M,n) & is_eclose_n(M,A,n,eclosen) & l \<in> eclosen"
+       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
@@ -607,10 +607,10 @@
    "M(A) \<Longrightarrow> iterates_replacement(M, big_union(M), A)"
   and eclose_replacement2:
    "M(A) \<Longrightarrow> strong_replacement(M,
-         \<lambda>n y. n\<in>nat & is_iterates(M, big_union(M), A, n, y))"
+         \<lambda>n y. n\<in>nat \<and> is_iterates(M, big_union(M), A, n, y))"
 
 lemma (in M_eclose) eclose_replacement2':
-  "M(A) \<Longrightarrow> strong_replacement(M, \<lambda>n y. n\<in>nat & y = Union^n (A))"
+  "M(A) \<Longrightarrow> strong_replacement(M, \<lambda>n y. n\<in>nat \<and> y = Union^n (A))"
 apply (insert eclose_replacement2 [of A])
 apply (rule strong_replacement_cong [THEN iffD1])
 apply (rule conj_cong [OF iff_refl iterates_abs [of "big_union(M)"]])
@@ -653,14 +653,14 @@
   is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o" where
    "is_transrec(M,MH,a,z) \<equiv>
       \<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M].
-       upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) &
+       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(M,MH,a) \<equiv>
       \<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M].
-       upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) &
+       upair(M,a,a,sa) \<and> is_eclose(M,sa,esa) \<and> membership(M,esa,mesa) \<and>
        wfrec_replacement(M,MH,mesa)"
 
 text\<open>The condition \<^term>\<open>Ord(i)\<close> lets us use the simpler
@@ -688,7 +688,7 @@
 lemma (in M_eclose) transrec_replacementI:
    "\<lbrakk>M(a);
       strong_replacement (M,
-                  \<lambda>x z. \<exists>y[M]. pair(M, x, y, z) &
+                  \<lambda>x z. \<exists>y[M]. pair(M, x, y, z) \<and>
                                is_wfrec(M,MH,Memrel(eclose({a})),x,y))\<rbrakk>
     \<Longrightarrow> transrec_replacement(M,MH,a)"
 by (simp add: transrec_replacement_def wfrec_replacement_def)
@@ -701,13 +701,13 @@
   is_length :: "[i=>o,i,i,i] => 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) & l \<notin> list_n &
-        successor(M,n,sn) & is_list_N(M,A,sn,list_sn) & l \<in> list_sn"
+        is_list_N(M,A,n,list_n) \<and> l \<notin> list_n \<and>
+        successor(M,n,sn) \<and> is_list_N(M,A,sn,list_sn) \<and> l \<in> list_sn"
 
 
 lemma (in M_datatypes) length_abs [simp]:
      "\<lbrakk>M(A); l \<in> list(A); n \<in> nat\<rbrakk> \<Longrightarrow> is_length(M,A,l,n) \<longleftrightarrow> n = length(l)"
-apply (subgoal_tac "M(l) & M(n)")
+apply (subgoal_tac "M(l) \<and> M(n)")
  prefer 2 apply (blast dest: transM)
 apply (simp add: is_length_def)
 apply (blast intro: list_imp_list_N nat_into_Ord list_N_imp_eq_length
@@ -748,7 +748,7 @@
 definition
   is_nth :: "[i=>o,i,i,i] => o" where
     "is_nth(M,n,l,Z) \<equiv>
-      \<exists>X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,Z)"
+      \<exists>X[M]. is_iterates(M, is_tl(M), l, n, X) \<and> is_hd(M,X,Z)"
 
 lemma (in M_datatypes) nth_abs [simp]:
      "\<lbrakk>M(A); n \<in> nat; l \<in> list(A); M(Z)\<rbrakk>
@@ -767,46 +767,46 @@
   is_Member :: "[i=>o,i,i,i] => 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) & is_Inl(M,p,u) & is_Inl(M,u,Z)"
+        \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) \<and> is_Inl(M,p,u) \<and> is_Inl(M,u,Z)"
 
 lemma (in M_trivial) Member_abs [simp]:
      "\<lbrakk>M(x); M(y); M(Z)\<rbrakk> \<Longrightarrow> is_Member(M,x,y,Z) \<longleftrightarrow> (Z = Member(x,y))"
 by (simp add: is_Member_def Member_def)
 
 lemma (in M_trivial) Member_in_M_iff [iff]:
-     "M(Member(x,y)) \<longleftrightarrow> M(x) & M(y)"
+     "M(Member(x,y)) \<longleftrightarrow> M(x) \<and> M(y)"
 by (simp add: Member_def)
 
 definition
   is_Equal :: "[i=>o,i,i,i] => 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) & is_Inr(M,p,u) & is_Inl(M,u,Z)"
+        \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) \<and> is_Inr(M,p,u) \<and> is_Inl(M,u,Z)"
 
 lemma (in M_trivial) Equal_abs [simp]:
      "\<lbrakk>M(x); M(y); M(Z)\<rbrakk> \<Longrightarrow> is_Equal(M,x,y,Z) \<longleftrightarrow> (Z = Equal(x,y))"
 by (simp add: is_Equal_def Equal_def)
 
-lemma (in M_trivial) Equal_in_M_iff [iff]: "M(Equal(x,y)) \<longleftrightarrow> M(x) & M(y)"
+lemma (in M_trivial) Equal_in_M_iff [iff]: "M(Equal(x,y)) \<longleftrightarrow> M(x) \<and> M(y)"
 by (simp add: Equal_def)
 
 definition
   is_Nand :: "[i=>o,i,i,i] => 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) & is_Inl(M,p,u) & is_Inr(M,u,Z)"
+        \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) \<and> is_Inl(M,p,u) \<and> is_Inr(M,u,Z)"
 
 lemma (in M_trivial) Nand_abs [simp]:
      "\<lbrakk>M(x); M(y); M(Z)\<rbrakk> \<Longrightarrow> is_Nand(M,x,y,Z) \<longleftrightarrow> (Z = Nand(x,y))"
 by (simp add: is_Nand_def Nand_def)
 
-lemma (in M_trivial) Nand_in_M_iff [iff]: "M(Nand(x,y)) \<longleftrightarrow> M(x) & M(y)"
+lemma (in M_trivial) Nand_in_M_iff [iff]: "M(Nand(x,y)) \<longleftrightarrow> M(x) \<and> M(y)"
 by (simp add: Nand_def)
 
 definition
   is_Forall :: "[i=>o,i,i] => 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) & is_Inr(M,u,Z)"
+    "is_Forall(M,p,Z) \<equiv> \<exists>u[M]. is_Inr(M,p,u) \<and> is_Inr(M,u,Z)"
 
 lemma (in M_trivial) Forall_abs [simp]:
      "\<lbrakk>M(x); M(Z)\<rbrakk> \<Longrightarrow> is_Forall(M,x,Z) \<longleftrightarrow> (Z = Forall(x))"
@@ -857,13 +857,13 @@
   is_depth :: "[i=>o,i,i] => 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) & p \<notin> formula_n &
-        successor(M,n,sn) & is_formula_N(M,sn,formula_sn) & p \<in> formula_sn"
+        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"
 
 
 lemma (in M_datatypes) depth_abs [simp]:
      "\<lbrakk>p \<in> formula; n \<in> nat\<rbrakk> \<Longrightarrow> is_depth(M,p,n) \<longleftrightarrow> n = depth(p)"
-apply (subgoal_tac "M(p) & M(n)")
+apply (subgoal_tac "M(p) \<and> M(n)")
  prefer 2 apply (blast dest: transM)
 apply (simp add: is_depth_def)
 apply (blast intro: formula_imp_formula_N nat_into_Ord formula_N_imp_eq_depth
@@ -884,11 +884,11 @@
   \<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>
-                      is_Member(M,x,y,p) \<longrightarrow> is_a(x,y,z)) &
+                      is_Member(M,x,y,p) \<longrightarrow> is_a(x,y,z)) \<and>
       (\<forall>x[M]. \<forall>y[M]. finite_ordinal(M,x) \<longrightarrow> finite_ordinal(M,y) \<longrightarrow>
-                      is_Equal(M,x,y,p) \<longrightarrow> is_b(x,y,z)) &
+                      is_Equal(M,x,y,p) \<longrightarrow> is_b(x,y,z)) \<and>
       (\<forall>x[M]. \<forall>y[M]. mem_formula(M,x) \<longrightarrow> mem_formula(M,y) \<longrightarrow>
-                     is_Nand(M,x,y,p) \<longrightarrow> is_c(x,y,z)) &
+                     is_Nand(M,x,y,p) \<longrightarrow> is_c(x,y,z)) \<and>
       (\<forall>x[M]. mem_formula(M,x) \<longrightarrow> is_Forall(M,x,p) \<longrightarrow> is_d(x,z))"
 
 lemma (in M_datatypes) formula_case_abs [simp]:
@@ -917,8 +917,8 @@
   is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => 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) & is_depth(M,p,dp) &
-             successor(M,dp,i) & fun_apply(M,f,p,z) & is_transrec(M,MH,i,f)"
+      \<exists>dp[M]. \<exists>i[M]. \<exists>f[M]. finite_ordinal(M,dp) \<and> is_depth(M,p,dp) \<and>
+             successor(M,dp,i) \<and> fun_apply(M,f,p,z) \<and> is_transrec(M,MH,i,f)"
 
 
 text\<open>Sufficient conditions to relativize the instance of \<^term>\<open>formula_case\<close>
@@ -968,7 +968,7 @@
       and fr_lam_replace:
            "M(g) \<Longrightarrow>
             strong_replacement
-              (M, \<lambda>x y. x \<in> formula &
+              (M, \<lambda>x y. x \<in> formula \<and>
                   y = \<langle>x, formula_rec_case(a,b,c,d,g,x)\<rangle>)"
 
 lemma (in Formula_Rec) formula_rec_case_closed:
--- a/src/ZF/Constructible/Formula.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/Formula.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -97,7 +97,7 @@
 
 lemma sats_Nand_iff [simp]:
   "env \<in> list(A)
-   \<Longrightarrow> (sats(A, Nand(p,q), env)) \<longleftrightarrow> \<not> (sats(A,p,env) & sats(A,q,env))"
+   \<Longrightarrow> (sats(A, Nand(p,q), env)) \<longleftrightarrow> \<not> (sats(A,p,env) \<and> sats(A,q,env))"
 by (simp add: Bool.and_def Bool.not_def cond_def)
 
 lemma sats_Forall_iff [simp]:
@@ -116,7 +116,7 @@
 
 lemma sats_And_iff [simp]:
   "env \<in> list(A)
-   \<Longrightarrow> (sats(A, And(p,q), env)) \<longleftrightarrow> sats(A,p,env) & sats(A,q,env)"
+   \<Longrightarrow> (sats(A, And(p,q), env)) \<longleftrightarrow> sats(A,p,env) \<and> sats(A,q,env)"
 by (simp add: And_def)
 
 lemma sats_Or_iff [simp]:
@@ -159,7 +159,7 @@
 
 lemma conj_iff_sats:
       "\<lbrakk>P \<longleftrightarrow> sats(A,p,env); Q \<longleftrightarrow> sats(A,q,env); env \<in> list(A)\<rbrakk>
-       \<Longrightarrow> (P & Q) \<longleftrightarrow> sats(A, And(p,q), env)"
+       \<Longrightarrow> (P \<and> Q) \<longleftrightarrow> sats(A, And(p,q), env)"
 by (simp)
 
 lemma disj_iff_sats:
@@ -394,7 +394,7 @@
   DPow :: "i => i" where
   "DPow(A) \<equiv> {X \<in> Pow(A).
                \<exists>env \<in> list(A). \<exists>p \<in> formula.
-                 arity(p) \<le> succ(length(env)) &
+                 arity(p) \<le> succ(length(env)) \<and>
                  X = {x\<in>A. sats(A, p, Cons(x,env))}}"
 
 lemma DPowI:
@@ -411,9 +411,9 @@
 
 lemma DPowD:
   "X \<in> DPow(A)
-   \<Longrightarrow> X \<subseteq> A &
+   \<Longrightarrow> X \<subseteq> A \<and>
        (\<exists>env \<in> list(A).
-        \<exists>p \<in> formula. arity(p) \<le> succ(length(env)) &
+        \<exists>p \<in> formula. arity(p) \<le> succ(length(env)) \<and>
                       X = {x\<in>A. sats(A, p, Cons(x,env))})"
 by (simp add: DPow_def)
 
@@ -593,7 +593,7 @@
 
 definition
   L :: "i=>o" where \<comment> \<open>Kunen's definition VI 1.5, page 167\<close>
-  "L(x) \<equiv> \<exists>i. Ord(i) & x \<in> Lset(i)"
+  "L(x) \<equiv> \<exists>i. Ord(i) \<and> x \<in> Lset(i)"
 
 text\<open>NOT SUITABLE FOR REWRITING -- RECURSIVE!\<close>
 lemma Lset: "Lset(i) = (\<Union>j\<in>i. DPow(Lset(j)))"
@@ -841,7 +841,7 @@
 lemma L_I: "\<lbrakk>x \<in> Lset(i); Ord(i)\<rbrakk> \<Longrightarrow> L(x)"
 by (simp add: L_def, blast)
 
-lemma L_D: "L(x) \<Longrightarrow> \<exists>i. Ord(i) & x \<in> Lset(i)"
+lemma L_D: "L(x) \<Longrightarrow> \<exists>i. Ord(i) \<and> x \<in> Lset(i)"
 by (simp add: L_def)
 
 lemma Ord_lrank [simp]: "Ord(lrank(a))"
@@ -859,7 +859,7 @@
 text\<open>Kunen's VI 1.8.  The proof is much harder than the text would
 suggest.  For a start, it needs the previous lemma, which is proved by
 induction.\<close>
-lemma Lset_iff_lrank_lt: "Ord(i) \<Longrightarrow> x \<in> Lset(i) \<longleftrightarrow> L(x) & lrank(x) < i"
+lemma Lset_iff_lrank_lt: "Ord(i) \<Longrightarrow> x \<in> Lset(i) \<longleftrightarrow> L(x) \<and> lrank(x) < i"
 apply (simp add: L_def, auto)
  apply (blast intro: Lset_lrank_lt)
  apply (unfold lrank_def)
@@ -923,7 +923,7 @@
 
 text\<open>Every set of constructible sets is included in some \<^term>\<open>Lset\<close>\<close>
 lemma subset_Lset:
-     "(\<forall>x\<in>A. L(x)) \<Longrightarrow> \<exists>i. Ord(i) & A \<subseteq> Lset(i)"
+     "(\<forall>x\<in>A. L(x)) \<Longrightarrow> \<exists>i. Ord(i) \<and> A \<subseteq> Lset(i)"
 by (rule_tac x = "\<Union>x\<in>A. succ(lrank(x))" in exI, force)
 
 lemma subset_LsetE:
@@ -940,7 +940,7 @@
 by (auto intro: L_I iff: Lset_succ_lrank_iff)
 
 lemma LPow_in_Lset:
-     "\<lbrakk>X \<in> Lset(i); Ord(i)\<rbrakk> \<Longrightarrow> \<exists>j. Ord(j) & {y \<in> Pow(X). L(y)} \<in> Lset(j)"
+     "\<lbrakk>X \<in> Lset(i); Ord(i)\<rbrakk> \<Longrightarrow> \<exists>j. Ord(j) \<and> {y \<in> Pow(X). L(y)} \<in> Lset(j)"
 apply (rule_tac x="succ(\<Union>y \<in> Pow(X). succ(lrank(y)))" in exI)
 apply simp
 apply (rule LsetI [OF succI1])
@@ -985,7 +985,7 @@
 
 lemma exists_bigger_env:
   "\<lbrakk>p \<in> formula; 0 \<in> A; env \<in> list(A)\<rbrakk>
-   \<Longrightarrow> \<exists>env' \<in> list(A). arity(p) \<le> succ(length(env')) &
+   \<Longrightarrow> \<exists>env' \<in> list(A). arity(p) \<le> succ(length(env')) \<and>
               (\<forall>a\<in>A. sats(A,p,Cons(a,env')) \<longleftrightarrow> sats(A,p,Cons(a,env)))"
 apply (rule_tac x="env @ repeat(0,arity(p))" in bexI)
 apply (simp del: app_Cons add: app_Cons [symmetric]
--- a/src/ZF/Constructible/Internalize.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/Internalize.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -8,7 +8,7 @@
 
 subsubsection\<open>The Formula \<^term>\<open>is_Inl\<close>, Internalized\<close>
 
-(*  is_Inl(M,a,z) \<equiv> \<exists>zero[M]. empty(M,zero) & pair(M,zero,a,z) *)
+(*  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(a,z) \<equiv> Exists(And(empty_fm(0), pair_fm(0,succ(a),succ(z))))"
@@ -38,7 +38,7 @@
 
 subsubsection\<open>The Formula \<^term>\<open>is_Inr\<close>, Internalized\<close>
 
-(*  is_Inr(M,a,z) \<equiv> \<exists>n1[M]. number1(M,n1) & pair(M,n1,a,z) *)
+(*  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(a,z) \<equiv> Exists(And(number1_fm(0), pair_fm(0,succ(a),succ(z))))"
@@ -68,7 +68,7 @@
 
 subsubsection\<open>The Formula \<^term>\<open>is_Nil\<close>, Internalized\<close>
 
-(* is_Nil(M,xs) \<equiv> \<exists>zero[M]. empty(M,zero) & is_Inl(M,zero,xs) *)
+(* is_Nil(M,xs) \<equiv> \<exists>zero[M]. empty(M,zero) \<and> is_Inl(M,zero,xs) *)
 
 definition
   Nil_fm :: "i=>i" where
@@ -98,7 +98,7 @@
 subsubsection\<open>The Formula \<^term>\<open>is_Cons\<close>, Internalized\<close>
 
 
-(*  "is_Cons(M,a,l,Z) \<equiv> \<exists>p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)" *)
+(*  "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(a,l,Z) \<equiv>
@@ -163,8 +163,8 @@
 subsubsection\<open>The Formula \<^term>\<open>is_hd\<close>, Internalized\<close>
 
 (*   "is_hd(M,xs,H) \<equiv> 
-       (is_Nil(M,xs) \<longrightarrow> empty(M,H)) &
-       (\<forall>x[M]. \<forall>l[M]. \<not> is_Cons(M,x,l,xs) | H=x) &
+       (is_Nil(M,xs) \<longrightarrow> empty(M,H)) \<and>
+       (\<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
@@ -200,8 +200,8 @@
 subsubsection\<open>The Formula \<^term>\<open>is_tl\<close>, Internalized\<close>
 
 (*     "is_tl(M,xs,T) \<equiv>
-       (is_Nil(M,xs) \<longrightarrow> T=xs) &
-       (\<forall>x[M]. \<forall>l[M]. \<not> is_Cons(M,x,l,xs) | T=l) &
+       (is_Nil(M,xs) \<longrightarrow> T=xs) \<and>
+       (\<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
@@ -237,7 +237,7 @@
 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(M,P,z) \<equiv> (P & number1(M,z)) | (\<not>P & empty(M,z))" *)
+   "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
@@ -282,7 +282,7 @@
 (* is_lambda :: "[i=>o, i, [i,i]=>o, i] => 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 & pair(M,u,v,p) & is_b(u,v))" *)
+        (\<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(p,A,z) \<equiv> 
@@ -322,7 +322,7 @@
 subsubsection\<open>The Operator \<^term>\<open>is_Member\<close>, Internalized\<close>
 
 (*    "is_Member(M,x,y,Z) \<equiv>
-        \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)" *)
+        \<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(x,y,Z) \<equiv>
@@ -355,7 +355,7 @@
 subsubsection\<open>The Operator \<^term>\<open>is_Equal\<close>, Internalized\<close>
 
 (*    "is_Equal(M,x,y,Z) \<equiv>
-        \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)" *)
+        \<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(x,y,Z) \<equiv>
@@ -388,7 +388,7 @@
 subsubsection\<open>The Operator \<^term>\<open>is_Nand\<close>, Internalized\<close>
 
 (*    "is_Nand(M,x,y,Z) \<equiv>
-        \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)" *)
+        \<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(x,y,Z) \<equiv>
@@ -420,7 +420,7 @@
 
 subsubsection\<open>The Operator \<^term>\<open>is_Forall\<close>, Internalized\<close>
 
-(* "is_Forall(M,p,Z) \<equiv> \<exists>u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)" *)
+(* "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(x,Z) \<equiv>
@@ -452,8 +452,8 @@
 
 subsubsection\<open>The Operator \<^term>\<open>is_and\<close>, Internalized\<close>
 
-(* is_and(M,a,b,z) \<equiv> (number1(M,a)  & z=b) | 
-                       (\<not>number1(M,a) & empty(M,z)) *)
+(* 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(a,b,z) \<equiv>
@@ -486,8 +486,8 @@
 
 subsubsection\<open>The Operator \<^term>\<open>is_or\<close>, Internalized\<close>
 
-(* is_or(M,a,b,z) \<equiv> (number1(M,a)  & number1(M,z)) | 
-                     (\<not>number1(M,a) & z=b) *)
+(* is_or(M,a,b,z) \<equiv> (number1(M,a)  \<and> number1(M,z)) | 
+                     (\<not>number1(M,a) \<and> z=b) *)
 
 definition
   or_fm :: "[i,i,i]=>i" where
@@ -522,8 +522,8 @@
 
 subsubsection\<open>The Operator \<^term>\<open>is_not\<close>, Internalized\<close>
 
-(* is_not(M,a,z) \<equiv> (number1(M,a)  & empty(M,z)) | 
-                     (\<not>number1(M,a) & number1(M,z)) *)
+(* 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(a,z) \<equiv>
@@ -568,10 +568,10 @@
    "M_is_recfun(M,MH,r,a,f) \<longleftrightarrow>
     (\<forall>z[M]. z \<in> f \<longleftrightarrow> 
      (\<exists>x[M]. \<exists>f_r_sx[M]. \<exists>y[M]. 
-             MH(x, f_r_sx, y) & pair(M,x,y,z) &
+             MH(x, f_r_sx, y) \<and> pair(M,x,y,z) \<and>
              (\<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. 
-                pair(M,x,a,xa) & upair(M,x,x,sx) &
-               pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) &
+                pair(M,x,a,xa) \<and> upair(M,x,x,sx) \<and>
+               pre_image(M,r,sx,r_sx) \<and> restriction(M,f,r_sx,f_r_sx) \<and>
                xa \<in> r)))"
 apply (simp add: M_is_recfun_def)
 apply (rule rall_cong, blast) 
@@ -583,10 +583,10 @@
      \<forall>z[M]. z \<in> f \<longleftrightarrow>
                2      1           0
 new def     (\<exists>x[M]. \<exists>f_r_sx[M]. \<exists>y[M]. 
-             MH(x, f_r_sx, y) & pair(M,x,y,z) &
+             MH(x, f_r_sx, y) \<and> pair(M,x,y,z) \<and>
              (\<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. 
-                pair(M,x,a,xa) & upair(M,x,x,sx) &
-               pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) &
+                pair(M,x,a,xa) \<and> upair(M,x,x,sx) \<and>
+               pre_image(M,r,sx,r_sx) \<and> restriction(M,f,r_sx,f_r_sx) \<and>
                xa \<in> r)"
 *)
 
@@ -653,7 +653,7 @@
 
 (* is_wfrec :: "[i=>o, i, [i,i,i]=>o, i, i] => o"
     "is_wfrec(M,MH,r,a,z) \<equiv> 
-      \<exists>f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,z)" *)
+      \<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(p,r,a,z) \<equiv> 
@@ -716,7 +716,7 @@
 definition
   cartprod_fm :: "[i,i,i]=>i" where
 (* "cartprod(M,A,B,z) \<equiv>
-        \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))" *)
+        \<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>
        Forall(Iff(Member(0,succ(z)),
                   Exists(And(Member(0,succ(succ(A))),
@@ -752,8 +752,8 @@
 (* "is_sum(M,A,B,Z) \<equiv>
        \<exists>A0[M]. \<exists>n1[M]. \<exists>s1[M]. \<exists>B1[M].
          3      2       1        0
-       number1(M,n1) & cartprod(M,n1,A,A0) & upair(M,n1,n1,s1) &
-       cartprod(M,s1,B,B1) & union(M,A0,B1,Z)"  *)
+       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(A,B,Z) \<equiv>
@@ -824,8 +824,8 @@
 
 (* is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o"
     "is_nat_case(M, a, is_b, k, z) \<equiv>
-       (empty(M,k) \<longrightarrow> z=a) &
-       (\<forall>m[M]. successor(M,m,k) \<longrightarrow> is_b(m,z)) &
+       (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
@@ -882,7 +882,7 @@
 
 (*  iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => 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) & isF(gm,u),
+        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
@@ -947,7 +947,7 @@
       \<^term>\<open>p\<close> is enclosed by 9 (??) quantifiers.\<close>
 
 (*    "is_iterates(M,isF,v,n,Z) \<equiv> 
-      \<exists>sn[M]. \<exists>msn[M]. successor(M,n,sn) & membership(M,sn,msn) &
+      \<exists>sn[M]. \<exists>msn[M]. successor(M,n,sn) \<and> membership(M,sn,msn) \<and>
        1       0       is_wfrec(M, iterates_MH(M,isF,v), msn, n, Z)"*)
 
 definition
@@ -1057,7 +1057,7 @@
 
 (* mem_eclose(M,A,l) \<equiv> 
       \<exists>n[M]. \<exists>eclosen[M]. 
-       finite_ordinal(M,n) & is_eclose_n(M,A,n,eclosen) & l \<in> eclosen *)
+       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(x,y) \<equiv>
@@ -1125,7 +1125,7 @@
   list_functor_fm :: "[i,i,i]=>i" where
 (* "is_list_functor(M,A,X,Z) \<equiv>
         \<exists>n1[M]. \<exists>AX[M].
-         number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)" *)
+         number1(M,n1) \<and> cartprod(M,A,X,AX) \<and> is_sum(M,n1,AX,Z)" *)
     "list_functor_fm(A,X,Z) \<equiv>
        Exists(Exists(
         And(number1_fm(1),
@@ -1159,7 +1159,7 @@
 subsubsection\<open>The Formula \<^term>\<open>is_list_N\<close>, Internalized\<close>
 
 (* "is_list_N(M,A,n,Z) \<equiv> 
-      \<exists>zero[M]. empty(M,zero) & 
+      \<exists>zero[M]. empty(M,zero) \<and> 
                 is_iterates(M, is_list_functor(M,A), zero, n, Z)" *)
 
 definition
@@ -1202,7 +1202,7 @@
 
 (* mem_list(M,A,l) \<equiv> 
       \<exists>n[M]. \<exists>listn[M]. 
-       finite_ordinal(M,n) & is_list_N(M,A,n,listn) & l \<in> listn *)
+       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(x,y) \<equiv>
@@ -1270,9 +1270,9 @@
 (*     "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
-          omega(M,nat') & cartprod(M,nat',nat',natnat) &
-          is_sum(M,natnat,natnat,natnatsum) &
-          cartprod(M,X,X,XX) & is_sum(M,XX,X,X3) &
+          omega(M,nat') \<and> cartprod(M,nat',nat',natnat) \<and>
+          is_sum(M,natnat,natnat,natnatsum) \<and>
+          cartprod(M,X,X,XX) \<and> is_sum(M,XX,X,X3) \<and>
           is_sum(M,natnatsum,X3,Z)" *)
     "formula_functor_fm(X,Z) \<equiv>
        Exists(Exists(Exists(Exists(Exists(
@@ -1310,7 +1310,7 @@
 subsubsection\<open>The Formula \<^term>\<open>is_formula_N\<close>, Internalized\<close>
 
 (*  "is_formula_N(M,n,Z) \<equiv> 
-      \<exists>zero[M]. empty(M,zero) & 
+      \<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
@@ -1352,7 +1352,7 @@
 
 (*  mem_formula(M,p) \<equiv> 
       \<exists>n[M]. \<exists>formn[M]. 
-       finite_ordinal(M,n) & is_formula_N(M,n,formn) & p \<in> formn *)
+       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(x) \<equiv>
@@ -1423,7 +1423,7 @@
    "is_transrec(M,MH,a,z) \<equiv> 
       \<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M]. 
        2       1         0
-       upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) &
+       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
--- a/src/ZF/Constructible/L_axioms.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/L_axioms.thy	Tue Sep 27 17:03:23 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) & Replace(X, %x y. Q(x,y) & L(y)) \<subseteq> Lset(j)"
-apply (rule_tac x="\<Union>y \<in> Replace(X, %x y. Q(x,y) & L(y)). succ(lrank(y))"
+      \<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))"
        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) & Replace(X, %x y. Q(x,y) & L(y)) \<subseteq> Y"
+      \<Longrightarrow> \<exists>Y. L(Y) \<and> Replace(X, %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 & P(x,u))\<rbrakk>
+    "\<lbrakk>\<forall>B[L]. separation(L, %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)
@@ -142,7 +142,7 @@
       terms become enormous!\<close>
 definition
   L_Reflects :: "[i=>o,[i,i]=>o] => prop"  (\<open>(3REFLECTS/ [_,/ _])\<close>) where
-    "REFLECTS[P,Q] \<equiv> (\<Or>Cl. Closed_Unbounded(Cl) &
+    "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))))"
 
 
@@ -258,7 +258,7 @@
 
 lemma ReflectsD:
      "\<lbrakk>REFLECTS[P,Q]; Ord(i)\<rbrakk>
-      \<Longrightarrow> \<exists>j. i<j & (\<forall>x \<in> Lset(j). P(x) \<longleftrightarrow> Q(j,x))"
+      \<Longrightarrow> \<exists>j. i<j \<and> (\<forall>x \<in> Lset(j). P(x) \<longleftrightarrow> Q(j,x))"
 apply (unfold L_Reflects_def Closed_Unbounded_def)
 apply (elim meta_exE, clarify)
 apply (blast dest!: UnboundedD)
@@ -512,7 +512,7 @@
 
 subsubsection\<open>The Number 1, Internalized\<close>
 
-(* "number1(M,a) \<equiv> (\<exists>x[M]. empty(M,x) & successor(M,x,a))" *)
+(* "number1(M,a) \<equiv> (\<exists>x[M]. empty(M,x) \<and> successor(M,x,a))" *)
 definition
   number1_fm :: "i=>i" where
     "number1_fm(a) \<equiv> Exists(And(empty_fm(0), succ_fm(0,succ(a))))"
@@ -542,7 +542,7 @@
 
 subsubsection\<open>Big Union, Internalized\<close>
 
-(*  "big_union(M,A,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>y[M]. y\<in>A & x \<in> y)" *)
+(*  "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(A,z) \<equiv>
@@ -696,7 +696,7 @@
 subsubsection\<open>Domain of a Relation, Internalized\<close>
 
 (* "is_domain(M,r,z) \<equiv>
-        \<forall>x[M]. (x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w))))" *)
+        \<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(r,z) \<equiv>
@@ -731,7 +731,7 @@
 subsubsection\<open>Range of a Relation, Internalized\<close>
 
 (* "is_range(M,r,z) \<equiv>
-        \<forall>y[M]. (y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w))))" *)
+        \<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(r,z) \<equiv>
@@ -766,8 +766,8 @@
 subsubsection\<open>Field of a Relation, Internalized\<close>
 
 (* "is_field(M,r,z) \<equiv>
-        \<exists>dr[M]. is_domain(M,r,dr) &
-            (\<exists>rr[M]. is_range(M,r,rr) & union(M,dr,rr,z))" *)
+        \<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(r,z) \<equiv>
@@ -803,7 +803,7 @@
 subsubsection\<open>Image under a Relation, Internalized\<close>
 
 (* "image(M,r,A,z) \<equiv>
-        \<forall>y[M]. (y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w))))" *)
+        \<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(r,A,z) \<equiv>
@@ -839,7 +839,7 @@
 subsubsection\<open>Pre-Image under a Relation, Internalized\<close>
 
 (* "pre_image(M,r,A,z) \<equiv>
-        \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))" *)
+        \<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(r,A,z) \<equiv>
@@ -876,7 +876,7 @@
 
 (* "fun_apply(M,f,x,y) \<equiv>
         (\<exists>xs[M]. \<exists>fxs[M].
-         upair(M,x,x,xs) & image(M,f,xs,fxs) & big_union(M,fxs,y))" *)
+         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(f,x,y) \<equiv>
@@ -981,7 +981,7 @@
 subsubsection\<open>Typed Functions, Internalized\<close>
 
 (* "typed_function(M,A,B,r) \<equiv>
-        is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) &
+        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
@@ -1042,8 +1042,8 @@
 (* "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].
-                pair(M,x,z,p) & pair(M,x,y,xy) & pair(M,y,z,yz) &
-                xy \<in> s & yz \<in> r)" *)
+                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(r,s,t) \<equiv>
@@ -1081,7 +1081,7 @@
 subsubsection\<open>Injections, Internalized\<close>
 
 (* "injection(M,A,B,f) \<equiv>
-        typed_function(M,A,B,f) &
+        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
@@ -1123,8 +1123,8 @@
 
 (*  surjection :: "[i=>o,i,i,i] => o"
     "surjection(M,A,B,f) \<equiv>
-        typed_function(M,A,B,f) &
-        (\<forall>y[M]. y\<in>B \<longrightarrow> (\<exists>x[M]. x\<in>A & fun_apply(M,f,x,y)))" *)
+        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(A,B,f) \<equiv>
@@ -1161,7 +1161,7 @@
 subsubsection\<open>Bijections, Internalized\<close>
 
 (*   bijection :: "[i=>o,i,i,i] => o"
-    "bijection(M,A,B,f) \<equiv> injection(M,A,B,f) & surjection(M,A,B,f)" *)
+    "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(A,B,f) \<equiv> And(injection_fm(A,B,f), surjection_fm(A,B,f))"
@@ -1194,7 +1194,7 @@
 
 
 (* "restriction(M,r,A,z) \<equiv>
-        \<forall>x[M]. x \<in> z \<longleftrightarrow> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))" *)
+        \<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(r,A,z) \<equiv>
@@ -1230,7 +1230,7 @@
 
 (*  order_isomorphism :: "[i=>o,i,i,i,i,i] => o"
    "order_isomorphism(M,A,r,B,s,f) \<equiv>
-        bijection(M,A,B,f) &
+        bijection(M,A,B,f) \<and>
         (\<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow>
           (\<forall>p[M]. \<forall>fx[M]. \<forall>fy[M]. \<forall>q[M].
             pair(M,x,y,p) \<longrightarrow> fun_apply(M,f,x,fx) \<longrightarrow> fun_apply(M,f,y,fy) \<longrightarrow>
@@ -1282,8 +1282,8 @@
 text\<open>A limit ordinal is a non-empty, successor-closed ordinal\<close>
 
 (* "limit_ordinal(M,a) \<equiv>
-        ordinal(M,a) & \<not> empty(M,a) &
-        (\<forall>x[M]. x\<in>a \<longrightarrow> (\<exists>y[M]. y\<in>a & successor(M,x,y)))" *)
+        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
   limit_ordinal_fm :: "i=>i" where
@@ -1320,7 +1320,7 @@
 subsubsection\<open>Finite Ordinals: The Predicate ``Is A Natural Number''\<close>
 
 (*     "finite_ordinal(M,a) \<equiv> 
-        ordinal(M,a) & \<not> limit_ordinal(M,a) & 
+        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
@@ -1355,7 +1355,7 @@
 
 subsubsection\<open>Omega: The Set of Natural Numbers\<close>
 
-(* omega(M,a) \<equiv> limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> \<not> limit_ordinal(M,x)) *)
+(* 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(x) \<equiv>
--- a/src/ZF/Constructible/Rank.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/Rank.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -13,25 +13,25 @@
 assumes well_ord_iso_separation:
      "\<lbrakk>M(A); M(f); M(r)\<rbrakk>
       \<Longrightarrow> separation (M, \<lambda>x. x\<in>A \<longrightarrow> (\<exists>y[M]. (\<exists>p[M].
-                     fun_apply(M,f,x,y) & pair(M,y,x,p) & p \<in> r)))"
+                     fun_apply(M,f,x,y) \<and> pair(M,y,x,p) \<and> p \<in> r)))"
   and obase_separation:
      \<comment> \<open>part of the order type formalization\<close>
      "\<lbrakk>M(A); M(r)\<rbrakk>
       \<Longrightarrow> separation(M, \<lambda>a. \<exists>x[M]. \<exists>g[M]. \<exists>mx[M]. \<exists>par[M].
-             ordinal(M,x) & membership(M,x,mx) & pred_set(M,A,a,r,par) &
+             ordinal(M,x) \<and> membership(M,x,mx) \<and> pred_set(M,A,a,r,par) \<and>
              order_isomorphism(M,par,r,x,mx,g))"
   and obase_equals_separation:
      "\<lbrakk>M(A); M(r)\<rbrakk>
       \<Longrightarrow> separation (M, \<lambda>x. x\<in>A \<longrightarrow> \<not>(\<exists>y[M]. \<exists>g[M].
-                              ordinal(M,y) & (\<exists>my[M]. \<exists>pxr[M].
-                              membership(M,y,my) & pred_set(M,A,x,r,pxr) &
+                              ordinal(M,y) \<and> (\<exists>my[M]. \<exists>pxr[M].
+                              membership(M,y,my) \<and> pred_set(M,A,x,r,pxr) \<and>
                               order_isomorphism(M,pxr,r,y,my,g))))"
   and omap_replacement:
      "\<lbrakk>M(A); M(r)\<rbrakk>
       \<Longrightarrow> strong_replacement(M,
              \<lambda>a z. \<exists>x[M]. \<exists>g[M]. \<exists>mx[M]. \<exists>par[M].
-             ordinal(M,x) & pair(M,a,x,z) & membership(M,x,mx) &
-             pred_set(M,A,a,r,par) & order_isomorphism(M,par,r,x,mx,g))"
+             ordinal(M,x) \<and> pair(M,a,x,z) \<and> membership(M,x,mx) \<and>
+             pred_set(M,A,a,r,par) \<and> order_isomorphism(M,par,r,x,mx,g))"
 
 
 text\<open>Inductive argument for Kunen's Lemma I 6.1, etc.
@@ -138,7 +138,7 @@
 definition  
   obase :: "[i=>o,i,i] => 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) & 
+   "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
@@ -146,12 +146,12 @@
     \<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> & Ord(x) & 
+         z \<in> f \<longleftrightarrow> (\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = <a,x> \<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(M,A,r,i) \<equiv> \<exists>f[M]. omap(M,A,r,f) & is_range(M,f,i)"
+   "otype(M,A,r,i) \<equiv> \<exists>f[M]. omap(M,A,r,f) \<and> is_range(M,f,i)"
 
 
 text\<open>Can also be proved with the premise \<^term>\<open>M(z)\<close> instead of
@@ -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> & Ord(x) & 
+          (\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = <a,x> \<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)
@@ -182,7 +182,7 @@
 lemma (in M_ordertype) otype_iff:
      "\<lbrakk>otype(M,A,r,i); M(A); M(r); M(i)\<rbrakk> 
       \<Longrightarrow> x \<in> i \<longleftrightarrow> 
-          (M(x) & Ord(x) & 
+          (M(x) \<and> Ord(x) \<and> 
            (\<exists>a\<in>A. \<exists>g[M]. g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x))))"
 apply (auto simp add: omap_iff otype_def)
  apply (blast intro: transM) 
@@ -372,7 +372,7 @@
 
 lemma (in M_ordertype) ordertype_exists:
      "\<lbrakk>wellordered(M,A,r); M(A); M(r)\<rbrakk>
-      \<Longrightarrow> \<exists>f[M]. (\<exists>i[M]. Ord(i) & f \<in> ord_iso(A, r, i, Memrel(i)))"
+      \<Longrightarrow> \<exists>f[M]. (\<exists>i[M]. Ord(i) \<and> f \<in> ord_iso(A, r, i, Memrel(i)))"
 apply (insert obase_exists [of A r] omap_exists [of A r] otype_exists [of A r], simp, clarify)
 apply (rename_tac i) 
 apply (subgoal_tac "Ord(i)", blast intro: omap_ord_iso_otype)
@@ -421,57 +421,57 @@
        (\<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) & union(M,i,gx,y),
+                     %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(M,i,j,k) \<equiv> 
-        (\<not> ordinal(M,i) & \<not> ordinal(M,j) & k=0) |
-        (\<not> ordinal(M,i) & ordinal(M,j) & k=j) |
-        (ordinal(M,i) & \<not> ordinal(M,j) & k=i) |
-        (ordinal(M,i) & ordinal(M,j) & 
-         (\<exists>f fj sj. M(f) & M(fj) & M(sj) & 
-                    successor(M,j,sj) & is_oadd_fun(M,i,sj,sj,f) & 
-                    fun_apply(M,f,j,fj) & fj = k))"
+        (\<not> ordinal(M,i) \<and> \<not> ordinal(M,j) \<and> k=0) |
+        (\<not> ordinal(M,i) \<and> ordinal(M,j) \<and> k=j) |
+        (ordinal(M,i) \<and> \<not> ordinal(M,j) \<and> k=i) |
+        (ordinal(M,i) \<and> ordinal(M,j) \<and> 
+         (\<exists>f fj sj. M(f) \<and> M(fj) \<and> M(sj) \<and> 
+                    successor(M,j,sj) \<and> is_oadd_fun(M,i,sj,sj,f) \<and> 
+                    fun_apply(M,f,j,fj) \<and> fj = k))"
 
 definition
  (*NEEDS RELATIVIZATION*)
   omult_eqns :: "[i,i,i,i] => o" where
     "omult_eqns(i,x,g,z) \<equiv>
-            Ord(x) & 
-            (x=0 \<longrightarrow> z=0) &
-            (\<forall>j. x = succ(j) \<longrightarrow> z = g`j ++ i) &
+            Ord(x) \<and> 
+            (x=0 \<longrightarrow> z=0) \<and>
+            (\<forall>j. x = succ(j) \<longrightarrow> z = g`j ++ i) \<and>
             (Limit(x) \<longrightarrow> z = \<Union>(g``x))"
 
 definition
   is_omult_fun :: "[i=>o,i,i,i] => o" where
     "is_omult_fun(M,i,j,f) \<equiv> 
-            (\<exists>df. M(df) & is_function(M,f) & 
-                  is_domain(M,f,df) & subset(M, j, df)) & 
+            (\<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(M,i,j,k) \<equiv> 
-        \<exists>f fj sj. M(f) & M(fj) & M(sj) & 
-                  successor(M,j,sj) & is_omult_fun(M,i,sj,f) & 
-                  fun_apply(M,f,j,fj) & fj = k"
+        \<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> 
+                  fun_apply(M,f,j,fj) \<and> fj = k"
 
 
 locale M_ord_arith = M_ordertype +
   assumes oadd_strong_replacement:
    "\<lbrakk>M(i); M(j)\<rbrakk> \<Longrightarrow>
     strong_replacement(M, 
-         \<lambda>x z. \<exists>y[M]. pair(M,x,y,z) & 
-                  (\<exists>f[M]. \<exists>fx[M]. is_oadd_fun(M,i,j,x,f) & 
-                           image(M,f,x,fx) & y = i \<union> fx))"
+         \<lambda>x z. \<exists>y[M]. pair(M,x,y,z) \<and> 
+                  (\<exists>f[M]. \<exists>fx[M]. is_oadd_fun(M,i,j,x,f) \<and> 
+                           image(M,f,x,fx) \<and> y = i \<union> fx))"
 
  and omult_strong_replacement':
    "\<lbrakk>M(i); M(j)\<rbrakk> \<Longrightarrow>
     strong_replacement(M, 
-         \<lambda>x z. \<exists>y[M]. z = <x,y> &
-             (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. THE z. omult_eqns(i,x,g,z),g) & 
+         \<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> 
              y = (THE z. omult_eqns(i, x, g, z))))" 
 
 
@@ -480,7 +480,7 @@
 lemma (in M_ord_arith) is_oadd_fun_iff:
    "\<lbrakk>a\<le>j; M(i); M(j); M(a); M(f)\<rbrakk> 
     \<Longrightarrow> is_oadd_fun(M,i,j,a,f) \<longleftrightarrow>
-        f \<in> a \<rightarrow> range(f) & (\<forall>x. M(x) \<longrightarrow> x < a \<longrightarrow> f`x = i \<union> f``x)"
+        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"]
@@ -494,8 +494,8 @@
 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> &
-                  (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. i \<union> g``x,g) & 
+            \<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> 
                   y = i \<union> g``x))" 
 apply (insert oadd_strong_replacement [of i j]) 
 apply (simp add: is_oadd_fun_def relation2_def
@@ -547,7 +547,7 @@
 
 lemma (in M_ord_arith) oadd_abs:
     "\<lbrakk>M(i); M(j); M(k)\<rbrakk> \<Longrightarrow> is_oadd(M,i,j,k) \<longleftrightarrow> k = i++j"
-apply (case_tac "Ord(i) & Ord(j)")
+apply (case_tac "Ord(i) \<and> Ord(j)")
  apply (simp add: Ord_oadd_abs)
 apply (auto simp add: is_oadd_def oadd_eq_if_raw_oadd)
 done
@@ -575,7 +575,7 @@
 lemma the_omult_eqns_0: "(THE z. omult_eqns(i,0,g,z)) = 0"
 by (simp add: omult_eqns_0)
 
-lemma omult_eqns_succ: "omult_eqns(i,succ(j),g,z) \<longleftrightarrow> Ord(j) & z = g`j ++ i"
+lemma omult_eqns_succ: "omult_eqns(i,succ(j),g,z) \<longleftrightarrow> Ord(j) \<and> z = g`j ++ i"
 by (simp add: omult_eqns_def)
 
 lemma the_omult_eqns_succ:
@@ -683,8 +683,8 @@
      "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)  & 
-                        M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f) &
+         (\<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>
                         is_range(M,f,y)))"
  and Ord_wfrank_separation:
      "M(r) \<Longrightarrow>
@@ -710,7 +710,7 @@
 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) & is_recfun(r^+, x, %x f. range(f), f) &
+                  pair(M,x,y,z) \<and> is_recfun(r^+, x, %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"])
@@ -730,7 +730,7 @@
   wellfoundedrank :: "[i=>o,i,i] => i" where
     "wellfoundedrank(M,r,A) \<equiv>
         {p. x\<in>A, \<exists>y[M]. \<exists>f[M]. 
-                       p = <x,y> & is_recfun(r^+, x, %x f. range(f), f) &
+                       p = <x,y> \<and> is_recfun(r^+, x, %x f. range(f), f) \<and>
                        y = range(f)}"
 
 lemma (in M_wfrank) exists_wfrank:
@@ -803,7 +803,7 @@
 apply (simp add: Transset_def, clarify, simp)
 apply (rename_tac x i f u)
 apply (frule is_recfun_imp_in_r, assumption)
-apply (subgoal_tac "M(u) & M(i) & M(x)")
+apply (subgoal_tac "M(u) \<and> M(i) \<and> M(x)")
  prefer 2 apply (blast dest: transM, clarify)
 apply (rule_tac a=u in rangeI)
 apply (rule_tac x=u in ReplaceI)
@@ -886,7 +886,7 @@
          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)
-apply (subgoal_tac "a\<in>A & b\<in>A")
+apply (subgoal_tac "a\<in>A \<and> b\<in>A")
  prefer 2 apply blast
 apply (simp add: lt_def Ord_wellfoundedrank, clarify)
 apply (frule exists_wfrank [of concl: _ b], erule (1) transM, assumption)
@@ -909,7 +909,7 @@
 
 lemma (in M_wfrank) wellfounded_imp_subset_rvimage:
      "\<lbrakk>wellfounded(M,r); r \<subseteq> A*A; M(r); M(A)\<rbrakk>
-      \<Longrightarrow> \<exists>i f. Ord(i) & r \<subseteq> rvimage(A, f, Memrel(i))"
+      \<Longrightarrow> \<exists>i f. Ord(i) \<and> r \<subseteq> rvimage(A, f, Memrel(i))"
 apply (rule_tac x="range(wellfoundedrank(M,r,A))" in exI)
 apply (rule_tac x="wellfoundedrank(M,r,A)" in exI)
 apply (simp add: Ord_range_wellfoundedrank, clarify)
--- a/src/ZF/Constructible/Rank_Separation.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/Rank_Separation.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -18,15 +18,15 @@
 
 lemma well_ord_iso_Reflects:
   "REFLECTS[\<lambda>x. x\<in>A \<longrightarrow>
-                (\<exists>y[L]. \<exists>p[L]. fun_apply(L,f,x,y) & pair(L,y,x,p) & p \<in> r),
+                (\<exists>y[L]. \<exists>p[L]. fun_apply(L,f,x,y) \<and> pair(L,y,x,p) \<and> p \<in> r),
         \<lambda>i x. x\<in>A \<longrightarrow> (\<exists>y \<in> Lset(i). \<exists>p \<in> Lset(i).
-                fun_apply(##Lset(i),f,x,y) & pair(##Lset(i),y,x,p) & p \<in> r)]"
+                fun_apply(##Lset(i),f,x,y) \<and> pair(##Lset(i),y,x,p) \<and> p \<in> r)]"
 by (intro FOL_reflections function_reflections)
 
 lemma well_ord_iso_separation:
      "\<lbrakk>L(A); L(f); L(r)\<rbrakk>
       \<Longrightarrow> separation (L, \<lambda>x. x\<in>A \<longrightarrow> (\<exists>y[L]. (\<exists>p[L].
-                     fun_apply(L,f,x,y) & pair(L,y,x,p) & p \<in> r)))"
+                     fun_apply(L,f,x,y) \<and> pair(L,y,x,p) \<and> p \<in> r)))"
 apply (rule gen_separation_multi [OF well_ord_iso_Reflects, of "{A,f,r}"], 
        auto)
 apply (rule_tac env="[A,f,r]" in DPow_LsetI)
@@ -38,10 +38,10 @@
 
 lemma obase_reflects:
   "REFLECTS[\<lambda>a. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
-             ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) &
+             ordinal(L,x) \<and> membership(L,x,mx) \<and> pred_set(L,A,a,r,par) \<and>
              order_isomorphism(L,par,r,x,mx,g),
         \<lambda>i a. \<exists>x \<in> Lset(i). \<exists>g \<in> Lset(i). \<exists>mx \<in> Lset(i). \<exists>par \<in> Lset(i).
-             ordinal(##Lset(i),x) & membership(##Lset(i),x,mx) & pred_set(##Lset(i),A,a,r,par) &
+             ordinal(##Lset(i),x) \<and> membership(##Lset(i),x,mx) \<and> pred_set(##Lset(i),A,a,r,par) \<and>
              order_isomorphism(##Lset(i),par,r,x,mx,g)]"
 by (intro FOL_reflections function_reflections fun_plus_reflections)
 
@@ -49,7 +49,7 @@
      \<comment> \<open>part of the order type formalization\<close>
      "\<lbrakk>L(A); L(r)\<rbrakk>
       \<Longrightarrow> separation(L, \<lambda>a. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
-             ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) &
+             ordinal(L,x) \<and> membership(L,x,mx) \<and> pred_set(L,A,a,r,par) \<and>
              order_isomorphism(L,par,r,x,mx,g))"
 apply (rule gen_separation_multi [OF obase_reflects, of "{A,r}"], auto)
 apply (rule_tac env="[A,r]" in DPow_LsetI)
@@ -61,20 +61,20 @@
 
 lemma obase_equals_reflects:
   "REFLECTS[\<lambda>x. x\<in>A \<longrightarrow> \<not>(\<exists>y[L]. \<exists>g[L].
-                ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L].
-                membership(L,y,my) & pred_set(L,A,x,r,pxr) &
+                ordinal(L,y) \<and> (\<exists>my[L]. \<exists>pxr[L].
+                membership(L,y,my) \<and> pred_set(L,A,x,r,pxr) \<and>
                 order_isomorphism(L,pxr,r,y,my,g))),
         \<lambda>i x. x\<in>A \<longrightarrow> \<not>(\<exists>y \<in> Lset(i). \<exists>g \<in> Lset(i).
-                ordinal(##Lset(i),y) & (\<exists>my \<in> Lset(i). \<exists>pxr \<in> Lset(i).
-                membership(##Lset(i),y,my) & pred_set(##Lset(i),A,x,r,pxr) &
+                ordinal(##Lset(i),y) \<and> (\<exists>my \<in> Lset(i). \<exists>pxr \<in> Lset(i).
+                membership(##Lset(i),y,my) \<and> pred_set(##Lset(i),A,x,r,pxr) \<and>
                 order_isomorphism(##Lset(i),pxr,r,y,my,g)))]"
 by (intro FOL_reflections function_reflections fun_plus_reflections)
 
 lemma obase_equals_separation:
      "\<lbrakk>L(A); L(r)\<rbrakk>
       \<Longrightarrow> separation (L, \<lambda>x. x\<in>A \<longrightarrow> \<not>(\<exists>y[L]. \<exists>g[L].
-                              ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L].
-                              membership(L,y,my) & pred_set(L,A,x,r,pxr) &
+                              ordinal(L,y) \<and> (\<exists>my[L]. \<exists>pxr[L].
+                              membership(L,y,my) \<and> pred_set(L,A,x,r,pxr) \<and>
                               order_isomorphism(L,pxr,r,y,my,g))))"
 apply (rule gen_separation_multi [OF obase_equals_reflects, of "{A,r}"], auto)
 apply (rule_tac env="[A,r]" in DPow_LsetI)
@@ -85,13 +85,13 @@
 subsubsection\<open>Replacement for \<^term>\<open>omap\<close>\<close>
 
 lemma omap_reflects:
- "REFLECTS[\<lambda>z. \<exists>a[L]. a\<in>B & (\<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
-     ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) &
-     pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g)),
- \<lambda>i z. \<exists>a \<in> Lset(i). a\<in>B & (\<exists>x \<in> Lset(i). \<exists>g \<in> Lset(i). \<exists>mx \<in> Lset(i).
+ "REFLECTS[\<lambda>z. \<exists>a[L]. a\<in>B \<and> (\<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
+     ordinal(L,x) \<and> pair(L,a,x,z) \<and> membership(L,x,mx) \<and>
+     pred_set(L,A,a,r,par) \<and> order_isomorphism(L,par,r,x,mx,g)),
+ \<lambda>i z. \<exists>a \<in> Lset(i). a\<in>B \<and> (\<exists>x \<in> Lset(i). \<exists>g \<in> Lset(i). \<exists>mx \<in> Lset(i).
         \<exists>par \<in> Lset(i).
-         ordinal(##Lset(i),x) & pair(##Lset(i),a,x,z) &
-         membership(##Lset(i),x,mx) & pred_set(##Lset(i),A,a,r,par) &
+         ordinal(##Lset(i),x) \<and> pair(##Lset(i),a,x,z) \<and>
+         membership(##Lset(i),x,mx) \<and> pred_set(##Lset(i),A,a,r,par) \<and>
          order_isomorphism(##Lset(i),par,r,x,mx,g))]"
 by (intro FOL_reflections function_reflections fun_plus_reflections)
 
@@ -99,8 +99,8 @@
      "\<lbrakk>L(A); L(r)\<rbrakk>
       \<Longrightarrow> strong_replacement(L,
              \<lambda>a z. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
-             ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) &
-             pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g))"
+             ordinal(L,x) \<and> pair(L,a,x,z) \<and> membership(L,x,mx) \<and>
+             pred_set(L,A,a,r,par) \<and> order_isomorphism(L,par,r,x,mx,g))"
 apply (rule strong_replacementI)
 apply (rule_tac u="{A,r,B}" in gen_separation_multi [OF omap_reflects], auto)
 apply (rule_tac env="[A,B,r]" in DPow_LsetI)
@@ -153,15 +153,15 @@
 subsubsection\<open>Replacement for \<^term>\<open>wfrank\<close>\<close>
 
 lemma wfrank_replacement_Reflects:
- "REFLECTS[\<lambda>z. \<exists>x[L]. x \<in> A &
+ "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)  &
-                        M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) &
+         (\<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>
                         is_range(L,f,y))),
- \<lambda>i z. \<exists>x \<in> Lset(i). x \<in> A &
+ \<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)  &
-         M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y), rplus, x, f) &
+       (\<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>
          is_range(##Lset(i),f,y)))]"
 by (intro FOL_reflections function_reflections fun_plus_reflections
              is_recfun_reflection tran_closure_reflection)
@@ -170,8 +170,8 @@
      "L(r) \<Longrightarrow>
       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)  &
-                        M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) &
+         (\<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>
                         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 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/Rec_Separation.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -20,14 +20,14 @@
 
 (* "rtran_closure_mem(M,A,r,p) \<equiv>
       \<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M].
-       omega(M,nnat) & n\<in>nnat & successor(M,n,n') &
-       (\<exists>f[M]. typed_function(M,n',A,f) &
-        (\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) &
-          fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) &
+       omega(M,nnat) \<and> n\<in>nnat \<and> successor(M,n,n') \<and>
+       (\<exists>f[M]. typed_function(M,n',A,f) \<and>
+        (\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) \<and> empty(M,zero) \<and>
+          fun_apply(M,f,zero,x) \<and> fun_apply(M,f,n,y)) \<and>
         (\<forall>j[M]. j\<in>n \<longrightarrow>
           (\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M].
-            fun_apply(M,f,j,fj) & successor(M,j,sj) &
-            fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"*)
+            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(A,r,p) \<equiv>
@@ -120,7 +120,7 @@
 subsubsection\<open>Transitive Closure of a Relation, Internalized\<close>
 
 (*  "tran_closure(M,r,t) \<equiv>
-         \<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" *)
+         \<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(r,s) \<equiv>
@@ -155,9 +155,9 @@
 
 lemma wellfounded_trancl_reflects:
   "REFLECTS[\<lambda>x. \<exists>w[L]. \<exists>wx[L]. \<exists>rp[L].
-                 w \<in> Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \<in> rp,
+                 w \<in> Z \<and> pair(L,w,x,wx) \<and> tran_closure(L,r,rp) \<and> wx \<in> rp,
    \<lambda>i x. \<exists>w \<in> Lset(i). \<exists>wx \<in> Lset(i). \<exists>rp \<in> Lset(i).
-       w \<in> Z & pair(##Lset(i),w,x,wx) & tran_closure(##Lset(i),r,rp) &
+       w \<in> Z \<and> pair(##Lset(i),w,x,wx) \<and> tran_closure(##Lset(i),r,rp) \<and>
        wx \<in> rp]"
 by (intro FOL_reflections function_reflections fun_plus_reflections
           tran_closure_reflection)
@@ -166,7 +166,7 @@
          "\<lbrakk>L(r); L(Z)\<rbrakk> \<Longrightarrow>
           separation (L, \<lambda>x.
               \<exists>w[L]. \<exists>wx[L]. \<exists>rp[L].
-               w \<in> Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \<in> rp)"
+               w \<in> Z \<and> pair(L,w,x,wx) \<and> tran_closure(L,r,rp) \<and> wx \<in> rp)"
 apply (rule gen_separation_multi [OF wellfounded_trancl_reflects, of "{r,Z}"],
        auto)
 apply (rule_tac env="[r,Z]" in DPow_LsetI)
@@ -218,16 +218,16 @@
 
 lemma list_replacement2_Reflects:
  "REFLECTS
-   [\<lambda>x. \<exists>u[L]. u \<in> B & u \<in> nat &
+   [\<lambda>x. \<exists>u[L]. u \<in> B \<and> u \<in> nat \<and>
                 is_iterates(L, is_list_functor(L, A), 0, u, x),
-    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & u \<in> nat &
+    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> u \<in> nat \<and>
                is_iterates(##Lset(i), is_list_functor(##Lset(i), A), 0, u, x)]"
 by (intro FOL_reflections 
           is_iterates_reflection list_functor_reflection)
 
 lemma list_replacement2:
    "L(A) \<Longrightarrow> strong_replacement(L,
-         \<lambda>n y. n\<in>nat & is_iterates(L, is_list_functor(L,A), 0, n, y))"
+         \<lambda>n y. n\<in>nat \<and> is_iterates(L, is_list_functor(L,A), 0, n, y))"
 apply (rule strong_replacementI)
 apply (rule_tac u="{A,B,0,nat}" 
          in gen_separation_multi [OF list_replacement2_Reflects], 
@@ -245,9 +245,9 @@
 need to expand iterates_replacement and wfrec_replacement*)
 lemma formula_replacement1_Reflects:
  "REFLECTS
-   [\<lambda>x. \<exists>u[L]. u \<in> B & (\<exists>y[L]. pair(L,u,y,x) &
+   [\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and>
          is_wfrec(L, iterates_MH(L, is_formula_functor(L), 0), memsn, u, y)),
-    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) &
+    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) \<and>
          is_wfrec(##Lset(i),
                   iterates_MH(##Lset(i),
                           is_formula_functor(##Lset(i)), 0), memsn, u, y))]"
@@ -268,16 +268,16 @@
 
 lemma formula_replacement2_Reflects:
  "REFLECTS
-   [\<lambda>x. \<exists>u[L]. u \<in> B & u \<in> nat &
+   [\<lambda>x. \<exists>u[L]. u \<in> B \<and> u \<in> nat \<and>
                 is_iterates(L, is_formula_functor(L), 0, u, x),
-    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & u \<in> nat &
+    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> u \<in> nat \<and>
                is_iterates(##Lset(i), is_formula_functor(##Lset(i)), 0, u, x)]"
 by (intro FOL_reflections 
           is_iterates_reflection formula_functor_reflection)
 
 lemma formula_replacement2:
    "strong_replacement(L,
-         \<lambda>n y. n\<in>nat & is_iterates(L, is_formula_functor(L), 0, n, y))"
+         \<lambda>n y. n\<in>nat \<and> is_iterates(L, is_formula_functor(L), 0, n, y))"
 apply (rule strong_replacementI)
 apply (rule_tac u="{B,0,nat}" 
          in gen_separation_multi [OF formula_replacement2_Reflects], 
@@ -293,7 +293,7 @@
 subsubsection\<open>The Formula \<^term>\<open>is_nth\<close>, Internalized\<close>
 
 (* "is_nth(M,n,l,Z) \<equiv>
-      \<exists>X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,Z)" *)
+      \<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(n,l,Z) \<equiv> 
@@ -333,9 +333,9 @@
 need to expand iterates_replacement and wfrec_replacement*)
 lemma nth_replacement_Reflects:
  "REFLECTS
-   [\<lambda>x. \<exists>u[L]. u \<in> B & (\<exists>y[L]. pair(L,u,y,x) &
+   [\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and>
          is_wfrec(L, iterates_MH(L, is_tl(L), z), memsn, u, y)),
-    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) &
+    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) \<and>
          is_wfrec(##Lset(i),
                   iterates_MH(##Lset(i),
                           is_tl(##Lset(i)), z), memsn, u, y))]"
@@ -380,9 +380,9 @@
 
 lemma eclose_replacement1_Reflects:
  "REFLECTS
-   [\<lambda>x. \<exists>u[L]. u \<in> B & (\<exists>y[L]. pair(L,u,y,x) &
+   [\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and>
          is_wfrec(L, iterates_MH(L, big_union(L), A), memsn, u, y)),
-    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) &
+    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) \<and>
          is_wfrec(##Lset(i),
                   iterates_MH(##Lset(i), big_union(##Lset(i)), A),
                   memsn, u, y))]"
@@ -403,15 +403,15 @@
 
 lemma eclose_replacement2_Reflects:
  "REFLECTS
-   [\<lambda>x. \<exists>u[L]. u \<in> B & u \<in> nat &
+   [\<lambda>x. \<exists>u[L]. u \<in> B \<and> u \<in> nat \<and>
                 is_iterates(L, big_union(L), A, u, x),
-    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & u \<in> nat &
+    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> u \<in> nat \<and>
                is_iterates(##Lset(i), big_union(##Lset(i)), A, u, x)]"
 by (intro FOL_reflections function_reflections is_iterates_reflection)
 
 lemma eclose_replacement2:
    "L(A) \<Longrightarrow> strong_replacement(L,
-         \<lambda>n y. n\<in>nat & is_iterates(L, big_union(L), A, n, y))"
+         \<lambda>n y. n\<in>nat \<and> is_iterates(L, big_union(L), A, n, y))"
 apply (rule strong_replacementI)
 apply (rule_tac u="{A,B,nat}" 
          in gen_separation_multi [OF eclose_replacement2_Reflects],
--- a/src/ZF/Constructible/Reflection.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/Reflection.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -35,16 +35,16 @@
       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)"
-  defines "M(x) \<equiv> \<exists>a. Ord(a) & x \<in> Mset(a)"
-      and "Reflects(Cl,P,Q) \<equiv> Closed_Unbounded(Cl) &
+  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) & P(<y,z>)) \<longrightarrow>
+  defines "F0(P,y) \<equiv> \<mu> b. (\<exists>z. M(z) \<and> P(<y,z>)) \<longrightarrow>
                                (\<exists>z\<in>Mset(b). P(<y,z>))"
       and "FF(P)   \<equiv> \<lambda>a. \<Union>y\<in>Mset(a). F0(P,y)"
-      and "ClEx(P,a) \<equiv> Limit(a) & normalize(FF(P),a) = a"
+      and "ClEx(P,a) \<equiv> Limit(a) \<and> normalize(FF(P),a) = a"
 
 begin 
 
@@ -54,7 +54,7 @@
 text\<open>Awkward: we need a version of \<open>ClEx_def\<close> as an equality
       at the level of classes, which do not really exist\<close>
 lemma ClEx_eq:
-     "ClEx(P) \<equiv> \<lambda>a. Limit(a) & normalize(FF(P),a) = a"
+     "ClEx(P) \<equiv> \<lambda>a. Limit(a) \<and> normalize(FF(P),a) = a"
 by (simp add: ClEx_def [symmetric])
 
 
@@ -70,26 +70,26 @@
 
 theorem And_reflection [intro]:
   "\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk>
-      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) & P'(x),
-                                      \<lambda>a x. Q(a,x) & Q'(a,x))"
+      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> C'(a), \<lambda>x. P(x) \<and> P'(x),
+                                      \<lambda>a x. Q(a,x) \<and> Q'(a,x))"
   by (simp add: Reflects_def Closed_Unbounded_Int, blast)
 
 theorem Or_reflection [intro]:
      "\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk>
-      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) | P'(x),
-                                      \<lambda>a x. Q(a,x) | Q'(a,x))"
+      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> C'(a), \<lambda>x. P(x) \<or> P'(x),
+                                      \<lambda>a x. Q(a,x) \<or> Q'(a,x))"
 by (simp add: Reflects_def Closed_Unbounded_Int, blast)
 
 theorem Imp_reflection [intro]:
      "\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk>
-      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & C'(a),
+      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> C'(a),
                    \<lambda>x. P(x) \<longrightarrow> P'(x),
                    \<lambda>a x. Q(a,x) \<longrightarrow> Q'(a,x))"
 by (simp add: Reflects_def Closed_Unbounded_Int, blast)
 
 theorem Iff_reflection [intro]:
      "\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk>
-      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & C'(a),
+      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> C'(a),
                    \<lambda>x. P(x) \<longleftrightarrow> P'(x),
                    \<lambda>a x. Q(a,x) \<longleftrightarrow> Q'(a,x))"
 by (simp add: Reflects_def Closed_Unbounded_Int, blast)
@@ -97,33 +97,32 @@
 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>)"
-apply (unfold F0_def M_def, clarify)
-apply (rule LeastI2)
-  apply (blast intro: Mset_mono [THEN subsetD])
- apply (blast intro: lt_Ord2, blast)
-done
+  "\<lbrakk>y\<in>Mset(a); Ord(a); M(z); P(<y,z>)\<rbrakk> \<Longrightarrow> \<exists>z\<in>Mset(F0(P,y)). P(<y,z>)"
+  unfolding F0_def M_def
+  apply clarify
+  apply (rule LeastI2)
+    apply (blast intro: Mset_mono [THEN subsetD])
+   apply (blast intro: lt_Ord2, blast)
+  done
 
 lemma Ord_F0 [intro,simp]: "Ord(F0(P,y))"
-by (simp add: F0_def)
+  by (simp add: F0_def)
 
 lemma Ord_FF [intro,simp]: "Ord(FF(P,y))"
-by (simp add: FF_def)
+  by (simp add: FF_def)
 
 lemma cont_Ord_FF: "cont_Ord(FF(P))"
-apply (insert Mset_cont)
-apply (simp add: cont_Ord_def FF_def, blast)
-done
+  using Mset_cont by (simp add: cont_Ord_def FF_def, blast)
 
 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>)"
-apply (simp add: FF_def)
-apply (simp_all add: cont_Ord_Union [of concl: Mset]
-                     Mset_cont Mset_mono_le not_emptyI)
-apply (blast intro: F0_works)
-done
+  "\<lbrakk>M(z); y\<in>Mset(a); P(<y,z>); Ord(a)\<rbrakk> \<Longrightarrow> \<exists>z\<in>Mset(FF(P,a)). P(<y,z>)"
+  apply (simp add: FF_def)
+  apply (simp_all add: cont_Ord_Union [of concl: Mset]
+      Mset_cont Mset_mono_le not_emptyI)
+  apply (blast intro: F0_works)
+  done
 
 lemma FFN_works:
      "\<lbrakk>M(z); y\<in>Mset(a); P(<y,z>); Ord(a)\<rbrakk>
@@ -156,7 +155,7 @@
 
 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) & P(<y,z>)"
+      \<Longrightarrow> \<exists>z. M(z) \<and> P(<y,z>)"
 apply (simp add: ClEx_def M_def)
 apply (blast dest: Cl_reflects
              intro: Limit_is_Ord Pair_in_Mset)
@@ -165,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) & P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
+      \<Longrightarrow> (\<exists>z. M(z) \<and> P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
 by (blast intro: dest: ClEx_downward ClEx_upward)
 
 text\<open>...and it is closed and unbounded\<close>
@@ -187,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) & P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
+      \<Longrightarrow> (\<exists>z. M(z) \<and> P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
 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 +214,8 @@
 
 lemma Ex_reflection_0:
      "Reflects(Cl,P0,Q0)
-      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(P0,a),
-                   \<lambda>x. \<exists>z. M(z) & P0(<x,z>),
+      \<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>))"
 apply (simp add: Reflects_def)
 apply (intro conjI Closed_Unbounded_Int)
@@ -227,7 +226,7 @@
 
 lemma All_reflection_0:
      "Reflects(Cl,P0,Q0)
-      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x.\<not>P0(x), a),
+      \<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>))"
 apply (simp only: all_iff_not_ex_not ball_iff_not_bex_not)
@@ -237,15 +236,15 @@
 
 theorem Ex_reflection [intro]:
      "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x)))
-      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a),
-                   \<lambda>x. \<exists>z. M(z) & P(x,z),
+      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> ClEx(\<lambda>x. P(fst(x),snd(x)), a),
+                   \<lambda>x. \<exists>z. M(z) \<and> P(x,z),
                    \<lambda>a x. \<exists>z\<in>Mset(a). Q(a,x,z))"
 by (rule Ex_reflection_0 [of _ " \<lambda>x. P(fst(x),snd(x))"
                                "\<lambda>a x. Q(a,fst(x),snd(x))", simplified])
 
 theorem All_reflection [intro]:
      "Reflects(Cl,  \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x)))
-      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. \<not>P(fst(x),snd(x)), a),
+      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> ClEx(\<lambda>x. \<not>P(fst(x),snd(x)), a),
                    \<lambda>x. \<forall>z. M(z) \<longrightarrow> P(x,z),
                    \<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))"
 by (rule All_reflection_0 [of _ "\<lambda>x. P(fst(x),snd(x))"
@@ -255,14 +254,14 @@
 
 theorem Rex_reflection [intro]:
      "Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x)))
-      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a),
+      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> ClEx(\<lambda>x. P(fst(x),snd(x)), a),
                    \<lambda>x. \<exists>z[M]. P(x,z),
                    \<lambda>a x. \<exists>z\<in>Mset(a). Q(a,x,z))"
 by (unfold rex_def, blast)
 
 theorem Rall_reflection [intro]:
      "Reflects(Cl,  \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x)))
-      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. \<not>P(fst(x),snd(x)), a),
+      \<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> ClEx(\<lambda>x. \<not>P(fst(x),snd(x)), a),
                    \<lambda>x. \<forall>z[M]. P(x,z),
                    \<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))"
 by (unfold rall_def, blast)
@@ -278,7 +277,7 @@
 proof state.\<close>
 schematic_goal
      "Reflects(?Cl,
-               \<lambda>x. \<exists>y. M(y) & x \<in> y,
+               \<lambda>x. \<exists>y. M(y) \<and> x \<in> y,
                \<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)"
 by fast
 
@@ -286,8 +285,8 @@
 in the class of reflecting ordinals.  The \<^term>\<open>Ord(a)\<close> is redundant,
 though harmless.\<close>
 lemma
-     "Reflects(\<lambda>a. Ord(a) & ClEx(\<lambda>x. fst(x) \<in> snd(x), a),
-               \<lambda>x. \<exists>y. M(y) & x \<in> y,
+     "Reflects(\<lambda>a. Ord(a) \<and> ClEx(\<lambda>x. fst(x) \<in> snd(x), a),
+               \<lambda>x. \<exists>y. M(y) \<and> x \<in> y,
                \<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)"
 by fast
 
@@ -295,31 +294,31 @@
 text\<open>Example 2\<close>
 schematic_goal
      "Reflects(?Cl,
-               \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y),
+               \<lambda>x. \<exists>y. M(y) \<and> (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y),
                \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)"
 by fast
 
 text\<open>Example 2'.  We give the reflecting class explicitly.\<close>
 lemma
   "Reflects
-    (\<lambda>a. (Ord(a) &
-          ClEx(\<lambda>x. \<not> (snd(x) \<subseteq> fst(fst(x)) \<longrightarrow> snd(x) \<in> snd(fst(x))), a)) &
+    (\<lambda>a. (Ord(a) \<and>
+          ClEx(\<lambda>x. \<not> (snd(x) \<subseteq> fst(fst(x)) \<longrightarrow> snd(x) \<in> snd(fst(x))), a)) \<and>
           ClEx(\<lambda>x. \<forall>z. M(z) \<longrightarrow> z \<subseteq> fst(x) \<longrightarrow> z \<in> snd(x), a),
-            \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y),
+            \<lambda>x. \<exists>y. M(y) \<and> (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y),
             \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)"
 by fast
 
 text\<open>Example 2''.  We expand the subset relation.\<close>
 schematic_goal
   "Reflects(?Cl,
-        \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> (\<forall>w. M(w) \<longrightarrow> w\<in>z \<longrightarrow> w\<in>x) \<longrightarrow> z\<in>y),
+        \<lambda>x. \<exists>y. M(y) \<and> (\<forall>z. M(z) \<longrightarrow> (\<forall>w. M(w) \<longrightarrow> w\<in>z \<longrightarrow> w\<in>x) \<longrightarrow> z\<in>y),
         \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). (\<forall>w\<in>Mset(a). w\<in>z \<longrightarrow> w\<in>x) \<longrightarrow> z\<in>y)"
 by fast
 
 text\<open>Example 2'''.  Single-step version, to reveal the reflecting class.\<close>
 schematic_goal
      "Reflects(?Cl,
-               \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y),
+               \<lambda>x. \<exists>y. M(y) \<and> (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y),
                \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)"
 apply (rule Ex_reflection)
 txt\<open>
@@ -339,21 +338,21 @@
 if \<^term>\<open>P\<close> is quantifier-free, since it is not being relativized.\<close>
 schematic_goal
      "Reflects(?Cl,
-               \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<in> y \<longleftrightarrow> z \<in> x & P(z)),
-               \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<in> y \<longleftrightarrow> z \<in> x & P(z))"
+               \<lambda>x. \<exists>y. M(y) \<and> (\<forall>z. M(z) \<longrightarrow> z \<in> y \<longleftrightarrow> z \<in> x \<and> P(z)),
+               \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<in> y \<longleftrightarrow> z \<in> x \<and> P(z))"
 by fast
 
 text\<open>Example 3'\<close>
 schematic_goal
      "Reflects(?Cl,
-               \<lambda>x. \<exists>y. M(y) & y = Collect(x,P),
+               \<lambda>x. \<exists>y. M(y) \<and> y = Collect(x,P),
                \<lambda>a x. \<exists>y\<in>Mset(a). y = Collect(x,P))"
 by fast
 
 text\<open>Example 3''\<close>
 schematic_goal
      "Reflects(?Cl,
-               \<lambda>x. \<exists>y. M(y) & y = Replace(x,P),
+               \<lambda>x. \<exists>y. M(y) \<and> y = Replace(x,P),
                \<lambda>a x. \<exists>y\<in>Mset(a). y = Replace(x,P))"
 by fast
 
@@ -361,7 +360,7 @@
 to be relativized.\<close>
 schematic_goal
      "Reflects(?Cl,
-               \<lambda>A. 0\<notin>A \<longrightarrow> (\<exists>f. M(f) & f \<in> (\<Prod>X \<in> A. X)),
+               \<lambda>A. 0\<notin>A \<longrightarrow> (\<exists>f. M(f) \<and> f \<in> (\<Prod>X \<in> A. X)),
                \<lambda>a A. 0\<notin>A \<longrightarrow> (\<exists>f\<in>Mset(a). f \<in> (\<Prod>X \<in> A. X)))"
 by fast
 
--- a/src/ZF/Constructible/Relative.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/Relative.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -20,21 +20,21 @@
 
 definition
   upair :: "[i=>o,i,i,i] => o" where
-    "upair(M,a,b,z) \<equiv> a \<in> z & b \<in> z & (\<forall>x[M]. x\<in>z \<longrightarrow> x = a | x = b)"
+    "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(M,a,b,z) \<equiv> \<exists>x[M]. upair(M,a,a,x) &
-                     (\<exists>y[M]. upair(M,a,b,y) & upair(M,x,y,z))"
+    "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(M,a,b,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a | x \<in> b"
+    "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(M,a,b,z) \<equiv> \<exists>x[M]. upair(M,a,a,x) & union(M,x,b,z)"
+    "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
@@ -42,15 +42,15 @@
 
 definition
   number1 :: "[i=>o,i] => o" where
-    "number1(M,a) \<equiv> \<exists>x[M]. empty(M,x) & successor(M,x,a)"
+    "number1(M,a) \<equiv> \<exists>x[M]. empty(M,x) \<and> successor(M,x,a)"
 
 definition
   number2 :: "[i=>o,i] => o" where
-    "number2(M,a) \<equiv> \<exists>x[M]. number1(M,x) & successor(M,x,a)"
+    "number2(M,a) \<equiv> \<exists>x[M]. number1(M,x) \<and> successor(M,x,a)"
 
 definition
   number3 :: "[i=>o,i] => o" where
-    "number3(M,a) \<equiv> \<exists>x[M]. number2(M,x) & successor(M,x,a)"
+    "number3(M,a) \<equiv> \<exists>x[M]. number2(M,x) \<and> successor(M,x,a)"
 
 definition
   powerset :: "[i=>o,i,i] => o" where
@@ -58,84 +58,84 @@
 
 definition
   is_Collect :: "[i=>o,i,i=>o,i] => o" where
-    "is_Collect(M,A,P,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> A & P(x)"
+    "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(M,A,P,z) \<equiv> \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,u))"
+    "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(M,a,b,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a & x \<in> b"
+    "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(M,a,b,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a & x \<notin> b"
+    "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(M,A,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>y[M]. y\<in>A & x \<in> y)"
+    "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(M,A,z) \<equiv>
-             (A=0 \<longrightarrow> z=0) &
+             (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(M,A,B,z) \<equiv>
-        \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))"
+        \<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(M,A,B,Z) \<equiv>
        \<exists>A0[M]. \<exists>n1[M]. \<exists>s1[M]. \<exists>B1[M].
-       number1(M,n1) & cartprod(M,n1,A,A0) & upair(M,n1,n1,s1) &
-       cartprod(M,s1,B,B1) & union(M,A0,B1,Z)"
+       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(M,a,z) \<equiv> \<exists>zero[M]. empty(M,zero) & pair(M,zero,a,z)"
+    "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(M,a,z) \<equiv> \<exists>n1[M]. number1(M,n1) & pair(M,n1,a,z)"
+    "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(M,r,z) \<equiv>
         \<forall>x[M]. x \<in> z \<longleftrightarrow>
-             (\<exists>w[M]. w\<in>r & (\<exists>u[M]. \<exists>v[M]. pair(M,u,v,w) & pair(M,v,u,x)))"
+             (\<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(M,r,A,z) \<equiv>
-        \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))"
+        \<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(M,r,z) \<equiv>
-        \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w)))"
+        \<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(M,r,A,z) \<equiv>
-        \<forall>y[M]. y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w)))"
+        \<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
     \<comment> \<open>the cleaner
-      \<^term>\<open>\<exists>r'[M]. is_converse(M,r,r') & is_domain(M,r',z)\<close>
+      \<^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
         \<^term>\<open>M(converse(r))\<close>.\<close>
     "is_range(M,r,z) \<equiv>
-        \<forall>y[M]. y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w)))"
+        \<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(M,r,z) \<equiv>
-        \<exists>dr[M]. \<exists>rr[M]. is_domain(M,r,dr) & is_range(M,r,rr) &
+        \<exists>dr[M]. \<exists>rr[M]. is_domain(M,r,dr) \<and> is_range(M,r,rr) \<and>
                         union(M,dr,rr,z)"
 
 definition
@@ -153,12 +153,12 @@
   fun_apply :: "[i=>o,i,i,i] => o" where
     "fun_apply(M,f,x,y) \<equiv>
         (\<exists>xs[M]. \<exists>fxs[M].
-         upair(M,x,x,xs) & image(M,f,xs,fxs) & big_union(M,fxs,y))"
+         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(M,A,B,r) \<equiv>
-        is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) &
+        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
@@ -171,30 +171,30 @@
     "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].
-                pair(M,x,z,p) & pair(M,x,y,xy) & pair(M,y,z,yz) &
-                xy \<in> s & yz \<in> r)"
+                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
   injection :: "[i=>o,i,i,i] => o" where
     "injection(M,A,B,f) \<equiv>
-        typed_function(M,A,B,f) &
+        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(M,A,B,f) \<equiv>
-        typed_function(M,A,B,f) &
-        (\<forall>y[M]. y\<in>B \<longrightarrow> (\<exists>x[M]. x\<in>A & fun_apply(M,f,x,y)))"
+        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(M,A,B,f) \<equiv> injection(M,A,B,f) & surjection(M,A,B,f)"
+    "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(M,r,A,z) \<equiv>
-        \<forall>x[M]. x \<in> z \<longleftrightarrow> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))"
+        \<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
@@ -203,43 +203,43 @@
 definition
   ordinal :: "[i=>o,i] => o" where
      \<comment> \<open>an ordinal is a transitive set of transitive sets\<close>
-    "ordinal(M,a) \<equiv> transitive_set(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> transitive_set(M,x))"
+    "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
     \<comment> \<open>a limit ordinal is a non-empty, successor-closed ordinal\<close>
     "limit_ordinal(M,a) \<equiv>
-        ordinal(M,a) & \<not> empty(M,a) &
-        (\<forall>x[M]. x\<in>a \<longrightarrow> (\<exists>y[M]. y\<in>a & successor(M,x,y)))"
+        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
     \<comment> \<open>a successor ordinal is any ordinal that is neither empty nor limit\<close>
     "successor_ordinal(M,a) \<equiv>
-        ordinal(M,a) & \<not> empty(M,a) & \<not> limit_ordinal(M,a)"
+        ordinal(M,a) \<and> \<not> empty(M,a) \<and> \<not> limit_ordinal(M,a)"
 
 definition
   finite_ordinal :: "[i=>o,i] => 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) & \<not> limit_ordinal(M,a) &
+        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
     \<comment> \<open>omega is a limit ordinal none of whose elements are limit\<close>
-    "omega(M,a) \<equiv> limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> \<not> limit_ordinal(M,x))"
+    "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(M,z) \<equiv> empty(M,z) | (\<exists>m[M]. successor(M,m,z))"
+    "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(M, a, is_b, k, z) \<equiv>
-       (empty(M,k) \<longrightarrow> z=a) &
-       (\<forall>m[M]. successor(M,m,k) \<longrightarrow> is_b(m,z)) &
-       (is_quasinat(M,k) | empty(M,z))"
+       (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
@@ -301,7 +301,7 @@
         to \<open>M\<close>.  We do not have separation as a scheme; every instance
         that we need must be assumed (and later proved) separately.\<close>
     "separation(M,P) \<equiv>
-        \<forall>z[M]. \<exists>y[M]. \<forall>x[M]. x \<in> y \<longleftrightarrow> x \<in> z & P(x)"
+        \<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
@@ -318,24 +318,24 @@
 definition
   univalent :: "[i=>o, i, [i,i]=>o] => o" where
     "univalent(M,A,P) \<equiv>
-        \<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. \<forall>z[M]. P(x,y) & P(x,z) \<longrightarrow> y=z)"
+        \<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(M,P) \<equiv>
       \<forall>A[M]. univalent(M,A,P) \<longrightarrow>
-      (\<exists>Y[M]. \<forall>b[M]. (\<exists>x[M]. x\<in>A & P(x,b)) \<longrightarrow> b \<in> Y)"
+      (\<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(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 & P(x,b)))"
+      (\<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(M) \<equiv>
-        \<forall>x[M]. (\<exists>y[M]. y\<in>x) \<longrightarrow> (\<exists>y[M]. y\<in>x & \<not>(\<exists>z[M]. z\<in>x & z \<in> y))"
+        \<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))"
 
 
 subsection\<open>A trivial consistency proof for $V_\omega$\<close>
@@ -354,7 +354,7 @@
 by (blast intro: univ0_downwards_mem)
 
 lemma univ0_Bex_abs [simp]:
-     "A \<in> univ(0) \<Longrightarrow> (\<exists>x\<in>A. x \<in> univ(0) & P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(x))"
+     "A \<in> univ(0) \<Longrightarrow> (\<exists>x\<in>A. x \<in> univ(0) \<and> P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(x))"
 by (blast intro: univ0_downwards_mem)
 
 text\<open>Congruence rule for separation: can assume the variable is in \<open>M\<close>\<close>
@@ -373,7 +373,7 @@
 by (simp add: univalent_def)
 
 lemma univalent_conjI2 [intro,simp]:
-     "univalent(M,A,Q) \<Longrightarrow> univalent(M, A, \<lambda>x y. P(x,y) & Q(x,y))"
+     "univalent(M,A,Q) \<Longrightarrow> univalent(M, A, \<lambda>x y. P(x,y) \<and> Q(x,y))"
 by (simp add: univalent_def, blast)
 
 text\<open>Congruence rule for replacement\<close>
@@ -487,16 +487,16 @@
 
 lemma replacementD:
     "\<lbrakk>replacement(M,P); M(A);  univalent(M,A,P)\<rbrakk>
-     \<Longrightarrow> \<exists>Y[M]. (\<forall>b[M]. ((\<exists>x[M]. x\<in>A & P(x,b)) \<longrightarrow> b \<in> Y))"
+     \<Longrightarrow> \<exists>Y[M]. (\<forall>b[M]. ((\<exists>x[M]. x\<in>A \<and> P(x,b)) \<longrightarrow> b \<in> Y))"
 by (simp add: replacement_def)
 
 lemma strong_replacementD:
     "\<lbrakk>strong_replacement(M,P); M(A);  univalent(M,A,P)\<rbrakk>
-     \<Longrightarrow> \<exists>Y[M]. (\<forall>b[M]. (b \<in> Y \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,b))))"
+     \<Longrightarrow> \<exists>Y[M]. (\<forall>b[M]. (b \<in> Y \<longleftrightarrow> (\<exists>x[M]. x\<in>A \<and> P(x,b))))"
 by (simp add: strong_replacement_def)
 
 lemma separationD:
-    "\<lbrakk>separation(M,P); M(z)\<rbrakk> \<Longrightarrow> \<exists>y[M]. \<forall>x[M]. x \<in> y \<longleftrightarrow> x \<in> z & P(x)"
+    "\<lbrakk>separation(M,P); M(z)\<rbrakk> \<Longrightarrow> \<exists>y[M]. \<forall>x[M]. x \<in> y \<longleftrightarrow> x \<in> z \<and> P(x)"
 by (simp add: separation_def)
 
 
@@ -505,7 +505,7 @@
 definition
   order_isomorphism :: "[i=>o,i,i,i,i,i] => o" where
     "order_isomorphism(M,A,r,B,s,f) \<equiv>
-        bijection(M,A,B,f) &
+        bijection(M,A,B,f) \<and>
         (\<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow>
           (\<forall>p[M]. \<forall>fx[M]. \<forall>fy[M]. \<forall>q[M].
             pair(M,x,y,p) \<longrightarrow> fun_apply(M,f,x,fx) \<longrightarrow> fun_apply(M,f,y,fy) \<longrightarrow>
@@ -514,12 +514,12 @@
 definition
   pred_set :: "[i=>o,i,i,i,i] => o" where
     "pred_set(M,A,x,r,B) \<equiv>
-        \<forall>y[M]. y \<in> B \<longleftrightarrow> (\<exists>p[M]. p\<in>r & y \<in> A & pair(M,y,x,p))"
+        \<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(M,A,r) \<equiv>
-        \<forall>p[M]. p \<in> r \<longleftrightarrow> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>A & x\<in>y & pair(M,x,y,p)))"
+        \<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)))"
 
 
 subsection\<open>Introducing a Transitive Class Model\<close>
@@ -566,12 +566,12 @@
 by (blast intro: transM)
 
 lemma (in M_trans) rex_abs [simp]:
-     "M(A) \<Longrightarrow> (\<exists>x[M]. x\<in>A & P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(x))"
+     "M(A) \<Longrightarrow> (\<exists>x[M]. x\<in>A \<and> P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(x))"
 by (blast intro: transM)
 
 lemma (in M_trans) ball_iff_equiv:
      "M(A) \<Longrightarrow> (\<forall>x[M]. (x\<in>A \<longleftrightarrow> P(x))) \<longleftrightarrow>
-               (\<forall>x\<in>A. P(x)) & (\<forall>x. P(x) \<longrightarrow> M(x) \<longrightarrow> x\<in>A)"
+               (\<forall>x\<in>A. P(x)) \<and> (\<forall>x. P(x) \<longrightarrow> M(x) \<longrightarrow> x\<in>A)"
 by (blast intro: transM)
 
 text\<open>Simplifies proofs of equalities when there's an iff-equality
@@ -604,15 +604,15 @@
 done
 
 lemma (in M_trivial) upair_in_MI [intro!]:
-     " M(a) & M(b) \<Longrightarrow> M({a,b})"
+     " M(a) \<and> M(b) \<Longrightarrow> M({a,b})"
 by (insert upair_ax, simp add: upair_ax_def)
 
 lemma (in M_trans) upair_in_MD [dest!]:
-     "M({a,b}) \<Longrightarrow> M(a) & M(b)"
+     "M({a,b}) \<Longrightarrow> M(a) \<and> M(b)"
   by (blast intro: transM)
 
 lemma (in M_trivial) upair_in_M_iff [simp]:
-  "M({a,b}) \<longleftrightarrow> M(a) & M(b)"
+  "M({a,b}) \<longleftrightarrow> M(a) \<and> M(b)"
   by blast
 
 lemma (in M_trivial) singleton_in_MI [intro!]:
@@ -634,19 +634,19 @@
 done
 
 lemma (in M_trans) pair_in_MD [dest!]:
-     "M(<a,b>) \<Longrightarrow> M(a) & M(b)"
+     "M(<a,b>) \<Longrightarrow> M(a) \<and> M(b)"
   by (simp add: Pair_def, blast intro: transM)
 
 lemma (in M_trivial) pair_in_MI [intro!]:
-     "M(a) & M(b) \<Longrightarrow> M(<a,b>)"
+     "M(a) \<and> M(b) \<Longrightarrow> M(<a,b>)"
   by (simp add: Pair_def)
 
 lemma (in M_trivial) pair_in_M_iff [simp]:
-     "M(<a,b>) \<longleftrightarrow> M(a) & M(b)"
+     "M(<a,b>) \<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) & M(y)"
+     "\<lbrakk><x,y> \<in> A; M(A)\<rbrakk> \<Longrightarrow> M(x) \<and> M(y)"
   by (blast dest: transM)
 
 lemma (in M_trivial) cartprod_abs [simp]:
@@ -746,7 +746,7 @@
 lemma (in M_trans) univalent_Replace_iff:
      "\<lbrakk>M(A); univalent(M,A,P);
          \<And>x y. \<lbrakk>x\<in>A; P(x,y)\<rbrakk> \<Longrightarrow> M(y)\<rbrakk>
-      \<Longrightarrow> u \<in> Replace(A,P) \<longleftrightarrow> (\<exists>x. x\<in>A & P(x,u))"
+      \<Longrightarrow> u \<in> Replace(A,P) \<longleftrightarrow> (\<exists>x. x\<in>A \<and> P(x,u))"
   unfolding Replace_iff univalent_def
   by (blast dest: transM)
 
@@ -788,13 +788,13 @@
 apply (simp add: RepFun_def)
 done
 
-lemma Replace_conj_eq: "{y . x \<in> A, x\<in>A & y=f(x)} = {y . x\<in>A, y=f(x)}"
+lemma Replace_conj_eq: "{y . x \<in> A, x\<in>A \<and> y=f(x)} = {y . x\<in>A, y=f(x)}"
 by simp
 
 text\<open>Better than \<open>RepFun_closed\<close> when having the formula \<^term>\<open>x\<in>A\<close>
       makes relativization easier.\<close>
 lemma (in M_trans) RepFun_closed2:
-     "\<lbrakk>strong_replacement(M, \<lambda>x y. x\<in>A & y = f(x)); M(A); \<forall>x\<in>A. M(f(x))\<rbrakk>
+     "\<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)))"
 apply (simp add: RepFun_def)
 apply (frule strong_replacement_closed, assumption)
@@ -807,7 +807,7 @@
  is_lambda :: "[i=>o, i, [i,i]=>o, i] => 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 & pair(M,u,v,p) & is_b(u,v))"
+        (\<exists>u[M]. \<exists>v[M]. u\<in>A \<and> pair(M,u,v,p) \<and> is_b(u,v))"
 
 lemma (in M_trivial) lam_closed:
      "\<lbrakk>strong_replacement(M, \<lambda>x y. y = <x,b(x)>); M(A); \<forall>x\<in>A. M(b(x))\<rbrakk>
@@ -816,7 +816,7 @@
 
 text\<open>Better than \<open>lam_closed\<close>: has the formula \<^term>\<open>x\<in>A\<close>\<close>
 lemma (in M_trivial) lam_closed2:
-  "\<lbrakk>strong_replacement(M, \<lambda>x y. x\<in>A & y = \<langle>x, b(x)\<rangle>);
+  "\<lbrakk>strong_replacement(M, \<lambda>x y. x\<in>A \<and> y = \<langle>x, b(x)\<rangle>);
      M(A); \<forall>m[M]. m\<in>A \<longrightarrow> M(b(m))\<rbrakk> \<Longrightarrow> M(Lambda(A,b))"
 apply (simp add: lam_def)
 apply (blast intro: RepFun_closed2 dest: transM)
@@ -951,7 +951,7 @@
 done
 
 lemma (in M_trivial) successor_ordinal_abs [simp]:
-     "M(a) \<Longrightarrow> successor_ordinal(M,a) \<longleftrightarrow> Ord(a) & (\<exists>b[M]. a = succ(b))"
+     "M(a) \<Longrightarrow> successor_ordinal(M,a) \<longleftrightarrow> Ord(a) \<and> (\<exists>b[M]. a = succ(b))"
 apply (simp add: successor_ordinal_def, safe)
 apply (drule Ord_cases_disj, auto)
 done
@@ -1006,7 +1006,7 @@
   primrec
       "natnumber_aux(M,0) = (\<lambda>x\<in>nat. if empty(M,x) then 1 else 0)"
       "natnumber_aux(M,succ(n)) =
-           (\<lambda>x\<in>nat. if (\<exists>y[M]. natnumber_aux(M,n)`y=1 & successor(M,y,x))
+           (\<lambda>x\<in>nat. if (\<exists>y[M]. natnumber_aux(M,n)`y=1 \<and> successor(M,y,x))
                      then 1 else 0)"
 
   definition
@@ -1026,36 +1026,36 @@
      "M(B) \<Longrightarrow> separation(M, \<lambda>x. x \<notin> B)"
   and cartprod_separation:
      "\<lbrakk>M(A); M(B)\<rbrakk>
-      \<Longrightarrow> separation(M, \<lambda>z. \<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,z)))"
+      \<Longrightarrow> separation(M, \<lambda>z. \<exists>x[M]. x\<in>A \<and> (\<exists>y[M]. y\<in>B \<and> pair(M,x,y,z)))"
   and image_separation:
      "\<lbrakk>M(A); M(r)\<rbrakk>
-      \<Longrightarrow> separation(M, \<lambda>y. \<exists>p[M]. p\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,p)))"
+      \<Longrightarrow> separation(M, \<lambda>y. \<exists>p[M]. p\<in>r \<and> (\<exists>x[M]. x\<in>A \<and> pair(M,x,y,p)))"
   and converse_separation:
      "M(r) \<Longrightarrow> separation(M,
-         \<lambda>z. \<exists>p[M]. p\<in>r & (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,p) & pair(M,y,x,z)))"
+         \<lambda>z. \<exists>p[M]. p\<in>r \<and> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,p) \<and> pair(M,y,x,z)))"
   and restrict_separation:
-     "M(A) \<Longrightarrow> separation(M, \<lambda>z. \<exists>x[M]. x\<in>A & (\<exists>y[M]. pair(M,x,y,z)))"
+     "M(A) \<Longrightarrow> separation(M, \<lambda>z. \<exists>x[M]. x\<in>A \<and> (\<exists>y[M]. pair(M,x,y,z)))"
   and comp_separation:
      "\<lbrakk>M(r); M(s)\<rbrakk>
       \<Longrightarrow> separation(M, \<lambda>xz. \<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M].
-                  pair(M,x,z,xz) & pair(M,x,y,xy) & pair(M,y,z,yz) &
-                  xy\<in>s & yz\<in>r)"
+                  pair(M,x,z,xz) \<and> pair(M,x,y,xy) \<and> pair(M,y,z,yz) \<and>
+                  xy\<in>s \<and> yz\<in>r)"
   and pred_separation:
-     "\<lbrakk>M(r); M(x)\<rbrakk> \<Longrightarrow> separation(M, \<lambda>y. \<exists>p[M]. p\<in>r & pair(M,y,x,p))"
+     "\<lbrakk>M(r); M(x)\<rbrakk> \<Longrightarrow> separation(M, \<lambda>y. \<exists>p[M]. p\<in>r \<and> pair(M,y,x,p))"
   and Memrel_separation:
-     "separation(M, \<lambda>z. \<exists>x[M]. \<exists>y[M]. pair(M,x,y,z) & x \<in> y)"
+     "separation(M, \<lambda>z. \<exists>x[M]. \<exists>y[M]. pair(M,x,y,z) \<and> x \<in> y)"
   and funspace_succ_replacement:
      "M(n) \<Longrightarrow>
       strong_replacement(M, \<lambda>p z. \<exists>f[M]. \<exists>b[M]. \<exists>nb[M]. \<exists>cnbf[M].
-                pair(M,f,b,p) & pair(M,n,b,nb) & is_cons(M,nb,f,cnbf) &
+                pair(M,f,b,p) \<and> pair(M,n,b,nb) \<and> is_cons(M,nb,f,cnbf) \<and>
                 upair(M,cnbf,cnbf,z))"
   and is_recfun_separation:
      \<comment> \<open>for well-founded recursion: used to prove \<open>is_recfun_equal\<close>\<close>
      "\<lbrakk>M(r); M(f); M(g); M(a); M(b)\<rbrakk>
      \<Longrightarrow> separation(M,
             \<lambda>x. \<exists>xa[M]. \<exists>xb[M].
-                pair(M,x,a,xa) & xa \<in> r & pair(M,x,b,xb) & xb \<in> r &
-                (\<exists>fx[M]. \<exists>gx[M]. fun_apply(M,f,x,fx) & fun_apply(M,g,x,gx) &
+                pair(M,x,a,xa) \<and> xa \<in> r \<and> pair(M,x,b,xb) \<and> xb \<in> r \<and>
+                (\<exists>fx[M]. \<exists>gx[M]. fun_apply(M,f,x,fx) \<and> fun_apply(M,g,x,gx) \<and>
                                    fx \<noteq> gx))"
      and power_ax:         "power_ax(M)"
 
@@ -1075,7 +1075,7 @@
 lemma (in M_basic) cartprod_iff:
      "\<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) & powerset(M,p1,p2) &
+          (\<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>})"
 apply (simp add: Pair_def cartprod_def, safe)
 defer 1
@@ -1138,7 +1138,7 @@
      "M(r) \<Longrightarrow>
       converse(r) =
       {z \<in> \<Union>(\<Union>(r)) * \<Union>(\<Union>(r)).
-       \<exists>p\<in>r. \<exists>x[M]. \<exists>y[M]. p = \<langle>x,y\<rangle> & z = \<langle>y,x\<rangle>}"
+       \<exists>p\<in>r. \<exists>x[M]. \<exists>y[M]. p = \<langle>x,y\<rangle> \<and> z = \<langle>y,x\<rangle>}"
 apply (rule equalityI)
  prefer 2 apply (blast dest: transM, clarify, simp)
 apply (simp add: Pair_def)
@@ -1266,7 +1266,7 @@
      "\<lbrakk>M(r); M(s)\<rbrakk>
       \<Longrightarrow> r O s =
           {xz \<in> domain(s) * range(r).
-            \<exists>x[M]. \<exists>y[M]. \<exists>z[M]. xz = \<langle>x,z\<rangle> & \<langle>x,y\<rangle> \<in> s & \<langle>y,z\<rangle> \<in> r}"
+            \<exists>x[M]. \<exists>y[M]. \<exists>z[M]. xz = \<langle>x,z\<rangle> \<and> \<langle>x,y\<rangle> \<in> s \<and> \<langle>y,z\<rangle> \<in> r}"
 apply (simp add: comp_def)
 apply (rule equalityI)
  apply clarify
@@ -1342,13 +1342,13 @@
 subsubsection\<open>Some Facts About Separation Axioms\<close>
 
 lemma (in M_basic) separation_conj:
-     "\<lbrakk>separation(M,P); separation(M,Q)\<rbrakk> \<Longrightarrow> separation(M, \<lambda>z. P(z) & Q(z))"
+     "\<lbrakk>separation(M,P); separation(M,Q)\<rbrakk> \<Longrightarrow> separation(M, \<lambda>z. P(z) \<and> Q(z))"
 by (simp del: separation_closed
          add: separation_iff Collect_Int_Collect_eq [symmetric])
 
 (*???equalities*)
 lemma Collect_Un_Collect_eq:
-     "Collect(A,P) \<union> Collect(A,Q) = Collect(A, %x. P(x) | Q(x))"
+     "Collect(A,P) \<union> Collect(A,Q) = Collect(A, %x. P(x) \<or> Q(x))"
 by blast
 
 lemma Diff_Collect_eq:
@@ -1362,7 +1362,7 @@
 
 
 lemma (in M_basic) separation_disj:
-     "\<lbrakk>separation(M,P); separation(M,Q)\<rbrakk> \<Longrightarrow> separation(M, \<lambda>z. P(z) | Q(z))"
+     "\<lbrakk>separation(M,P); separation(M,Q)\<rbrakk> \<Longrightarrow> separation(M, \<lambda>z. P(z) \<or> Q(z))"
 by (simp del: separation_closed
          add: separation_iff Collect_Un_Collect_eq [symmetric])
 
@@ -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> & z = {cons(<n,b>, f)}}"
+      \<Union>{z. p \<in> (n->B)*B, \<exists>f[M]. \<exists>b[M]. p = <f,b> \<and> z = {cons(<n,b>, f)}}"
 apply (simp add: succ_fun_eq)
 apply (blast dest: transM)
 done
@@ -1430,22 +1430,22 @@
 
 definition
   is_bool_of_o :: "[i=>o, o, i] => o" where
-   "is_bool_of_o(M,P,z) \<equiv> (P & number1(M,z)) | (\<not>P & empty(M,z))"
+   "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(M,a,z) \<equiv> (number1(M,a)  & empty(M,z)) |
-                     (\<not>number1(M,a) & number1(M,z))"
+   "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(M,a,b,z) \<equiv> (number1(M,a)  & z=b) |
-                       (\<not>number1(M,a) & empty(M,z))"
+   "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(M,a,b,z) \<equiv> (number1(M,a)  & number1(M,z)) |
-                      (\<not>number1(M,a) & z=b)"
+   "is_or(M,a,b,z) \<equiv> (number1(M,a)  \<and> number1(M,z)) |
+                      (\<not>number1(M,a) \<and> z=b)"
 
 lemma (in M_trivial) bool_of_o_abs [simp]:
      "M(z) \<Longrightarrow> is_bool_of_o(M,P,z) \<longleftrightarrow> z = bool_of_o(P)"
@@ -1487,12 +1487,12 @@
 definition
   is_Nil :: "[i=>o, i] => o" where
      \<comment> \<open>because \<^prop>\<open>[] \<equiv> Inl(0)\<close>\<close>
-    "is_Nil(M,xs) \<equiv> \<exists>zero[M]. empty(M,zero) & is_Inl(M,zero,xs)"
+    "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
      \<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) & is_Inr(M,p,Z)"
+    "is_Cons(M,a,l,Z) \<equiv> \<exists>p[M]. pair(M,a,l,p) \<and> is_Inr(M,p,Z)"
 
 
 lemma (in M_trivial) Nil_in_M [intro,simp]: "M(Nil)"
@@ -1501,7 +1501,7 @@
 lemma (in M_trivial) Nil_abs [simp]: "M(Z) \<Longrightarrow> is_Nil(M,Z) \<longleftrightarrow> (Z = Nil)"
 by (simp add: is_Nil_def Nil_def)
 
-lemma (in M_trivial) Cons_in_M_iff [iff]: "M(Cons(a,l)) \<longleftrightarrow> M(a) & M(l)"
+lemma (in M_trivial) Cons_in_M_iff [iff]: "M(Cons(a,l)) \<longleftrightarrow> M(a) \<and> M(l)"
 by (simp add: Cons_def)
 
 lemma (in M_trivial) Cons_abs [simp]:
@@ -1511,11 +1511,11 @@
 
 definition
   quasilist :: "i => o" where
-    "quasilist(xs) \<equiv> xs=Nil | (\<exists>x l. xs = Cons(x,l))"
+    "quasilist(xs) \<equiv> xs=Nil \<or> (\<exists>x l. xs = Cons(x,l))"
 
 definition
   is_quasilist :: "[i=>o,i] => o" where
-    "is_quasilist(M,z) \<equiv> is_Nil(M,z) | (\<exists>x[M]. \<exists>l[M]. is_Cons(M,x,l,z))"
+    "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
@@ -1527,9 +1527,9 @@
   is_list_case :: "[i=>o, i, [i,i,i]=>o, i, i] => 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) &
-       (\<forall>x[M]. \<forall>l[M]. is_Cons(M,x,l,xs) \<longrightarrow> is_b(x,l,z)) &
-       (is_quasilist(M,xs) | empty(M,z))"
+       (is_Nil(M,xs) \<longrightarrow> z=a) \<and>
+       (\<forall>x[M]. \<forall>l[M]. is_Cons(M,x,l,xs) \<longrightarrow> is_b(x,l,z)) \<and>
+       (is_quasilist(M,xs) \<or> empty(M,z))"
 
 definition
   hd' :: "i => i" where
@@ -1546,17 +1546,17 @@
      \<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>
-       (is_Nil(M,xs) \<longrightarrow> empty(M,H)) &
-       (\<forall>x[M]. \<forall>l[M]. \<not> is_Cons(M,x,l,xs) | H=x) &
-       (is_quasilist(M,xs) | empty(M,H))"
+       (is_Nil(M,xs) \<longrightarrow> empty(M,H)) \<and>
+       (\<forall>x[M]. \<forall>l[M]. \<not> is_Cons(M,x,l,xs) \<or> H=x) \<and>
+       (is_quasilist(M,xs) \<or> empty(M,H))"
 
 definition
   is_tl :: "[i=>o,i,i] => 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) &
-       (\<forall>x[M]. \<forall>l[M]. \<not> is_Cons(M,x,l,xs) | T=l) &
-       (is_quasilist(M,xs) | empty(M,T))"
+       (is_Nil(M,xs) \<longrightarrow> T=xs) \<and>
+       (\<forall>x[M]. \<forall>l[M]. \<not> is_Cons(M,x,l,xs) \<or> T=l) \<and>
+       (is_quasilist(M,xs) \<or> empty(M,T))"
 
 subsubsection\<open>\<^term>\<open>quasilist\<close>: For Case-Splitting with \<^term>\<open>list_case'\<close>\<close>
 
--- a/src/ZF/Constructible/Satisfies_absolute.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/Satisfies_absolute.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -14,8 +14,8 @@
 (*    "is_depth(M,p,n) \<equiv> 
        \<exists>sn[M]. \<exists>formula_n[M]. \<exists>formula_sn[M]. 
          2          1                0
-        is_formula_N(M,n,formula_n) & p \<notin> formula_n &
-        successor(M,n,sn) & is_formula_N(M,sn,formula_sn) & p \<in> formula_sn" *)
+        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(p,n) \<equiv> 
@@ -60,10 +60,10 @@
 (* is_formula_case :: 
     "[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => 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)) &
-      (\<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)) &
+      (\<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>
       (\<forall>x[M]. \<forall>y[M]. x\<in>formula \<longrightarrow> y\<in>formula \<longrightarrow> 
-                     is_Nand(M,x,y,v) \<longrightarrow> is_c(x,y,z)) &
+                     is_Nand(M,x,y,v) \<longrightarrow> is_c(x,y,z)) \<and>
       (\<forall>x[M]. x\<in>formula \<longrightarrow> is_Forall(M,x,v) \<longrightarrow> is_d(x,z))" *)
 
 definition
@@ -179,8 +179,8 @@
    \<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]. 
-        finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) &
-        fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z)"
+        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)"
 
 lemma (in M_datatypes) is_depth_apply_abs [simp]:
      "\<lbrakk>M(h); p \<in> formula; M(z)\<rbrakk> 
@@ -206,7 +206,7 @@
              is_lambda(M, lA, 
                 \<lambda>env z. is_bool_of_o(M, 
                       \<exists>nx[M]. \<exists>ny[M]. 
-                       is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & nx \<in> ny, z),
+                       is_nth(M,x,env,nx) \<and> is_nth(M,y,env,ny) \<and> nx \<in> ny, z),
                 zz)"
 
 definition
@@ -222,7 +222,7 @@
     \<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, 
                 \<lambda>env z. is_bool_of_o(M, 
-                      \<exists>nx[M]. is_nth(M,x,env,nx) & is_nth(M,y,env,nx), z),
+                      \<exists>nx[M]. is_nth(M,x,env,nx) \<and> is_nth(M,y,env,nx), z),
                 zz)"
 
 definition 
@@ -234,9 +234,9 @@
    "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]. 
-                 (\<exists>rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) & 
-                 (\<exists>rq[M]. is_depth_apply(M,h,q,rq) & fun_apply(M,rq,env,hq)) & 
-                 (\<exists>pq[M]. is_and(M,hp,hq,pq) & is_not(M,pq,z)),
+                 (\<exists>rp[M]. is_depth_apply(M,h,p,rp) \<and> fun_apply(M,rp,env,hp)) \<and> 
+                 (\<exists>rq[M]. is_depth_apply(M,h,q,rq) \<and> fun_apply(M,rq,env,hq)) \<and> 
+                 (\<exists>pq[M]. is_and(M,hp,hq,pq) \<and> is_not(M,pq,z)),
                 zz)"
 
 definition
@@ -249,7 +249,7 @@
    "satisfies_is_d(M,A,h) \<equiv> 
     \<lambda>p zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, 
-                \<lambda>env z. \<exists>rp[M]. is_depth_apply(M,h,p,rp) & 
+                \<lambda>env z. \<exists>rp[M]. is_depth_apply(M,h,p,rp) \<and> 
                     is_bool_of_o(M, 
                            \<forall>x[M]. \<forall>xenv[M]. \<forall>hp[M]. 
                               x\<in>A \<longrightarrow> is_Cons(M,x,env,xenv) \<longrightarrow> 
@@ -294,36 +294,36 @@
     "\<lbrakk>M(A); x \<in> nat; y \<in> nat\<rbrakk>
      \<Longrightarrow> strong_replacement
          (M, \<lambda>env z. \<exists>bo[M]. \<exists>nx[M]. \<exists>ny[M]. 
-              env \<in> list(A) & is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & 
-              is_bool_of_o(M, nx \<in> ny, bo) &
+              env \<in> list(A) \<and> is_nth(M,x,env,nx) \<and> is_nth(M,y,env,ny) \<and> 
+              is_bool_of_o(M, nx \<in> ny, bo) \<and>
               pair(M, env, bo, z))"
  and
    Equal_replacement:
     "\<lbrakk>M(A); x \<in> nat; y \<in> nat\<rbrakk>
      \<Longrightarrow> strong_replacement
          (M, \<lambda>env z. \<exists>bo[M]. \<exists>nx[M]. \<exists>ny[M]. 
-              env \<in> list(A) & is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & 
-              is_bool_of_o(M, nx = ny, bo) &
+              env \<in> list(A) \<and> is_nth(M,x,env,nx) \<and> is_nth(M,y,env,ny) \<and> 
+              is_bool_of_o(M, nx = ny, bo) \<and>
               pair(M, env, bo, z))"
  and
    Nand_replacement:
     "\<lbrakk>M(A); M(rp); M(rq)\<rbrakk>
      \<Longrightarrow> strong_replacement
          (M, \<lambda>env z. \<exists>rpe[M]. \<exists>rqe[M]. \<exists>andpq[M]. \<exists>notpq[M]. 
-               fun_apply(M,rp,env,rpe) & fun_apply(M,rq,env,rqe) & 
-               is_and(M,rpe,rqe,andpq) & is_not(M,andpq,notpq) & 
-               env \<in> list(A) & pair(M, env, notpq, z))"
+               fun_apply(M,rp,env,rpe) \<and> fun_apply(M,rq,env,rqe) \<and> 
+               is_and(M,rpe,rqe,andpq) \<and> is_not(M,andpq,notpq) \<and> 
+               env \<in> list(A) \<and> pair(M, env, notpq, z))"
  and
   Forall_replacement:
    "\<lbrakk>M(A); M(rp)\<rbrakk>
     \<Longrightarrow> strong_replacement
         (M, \<lambda>env z. \<exists>bo[M]. 
-              env \<in> list(A) & 
+              env \<in> list(A) \<and> 
               is_bool_of_o (M, 
                             \<forall>a[M]. \<forall>co[M]. \<forall>rpco[M]. 
                                a\<in>A \<longrightarrow> is_Cons(M,a,env,co) \<longrightarrow>
                                fun_apply(M,rp,co,rpco) \<longrightarrow> number1(M, rpco), 
-                            bo) &
+                            bo) \<and>
               pair(M,env,bo,z))"
  and
   formula_rec_replacement: 
@@ -334,39 +334,39 @@
       \<comment> \<open>For the \<open>\<lambda>-abstraction\<close> in the \<^term>\<open>transrec\<close> body\<close>
    "\<lbrakk>M(g); M(A)\<rbrakk> \<Longrightarrow>
     strong_replacement (M, 
-       \<lambda>x y. mem_formula(M,x) &
+       \<lambda>x y. mem_formula(M,x) \<and>
              (\<exists>c[M]. is_formula_case(M, satisfies_is_a(M,A),
                                   satisfies_is_b(M,A),
                                   satisfies_is_c(M,A,g),
-                                  satisfies_is_d(M,A,g), x, c) &
+                                  satisfies_is_d(M,A,g), x, c) \<and>
              pair(M, x, c, y)))"
 
 
 lemma (in M_satisfies) Member_replacement':
     "\<lbrakk>M(A); x \<in> nat; y \<in> nat\<rbrakk>
      \<Longrightarrow> strong_replacement
-         (M, \<lambda>env z. env \<in> list(A) &
+         (M, \<lambda>env z. env \<in> list(A) \<and>
                      z = \<langle>env, bool_of_o(nth(x, env) \<in> nth(y, env))\<rangle>)"
 by (insert Member_replacement, simp) 
 
 lemma (in M_satisfies) Equal_replacement':
     "\<lbrakk>M(A); x \<in> nat; y \<in> nat\<rbrakk>
      \<Longrightarrow> strong_replacement
-         (M, \<lambda>env z. env \<in> list(A) &
+         (M, \<lambda>env z. env \<in> list(A) \<and>
                      z = \<langle>env, bool_of_o(nth(x, env) = nth(y, env))\<rangle>)"
 by (insert Equal_replacement, simp) 
 
 lemma (in M_satisfies) Nand_replacement':
     "\<lbrakk>M(A); M(rp); M(rq)\<rbrakk>
      \<Longrightarrow> strong_replacement
-         (M, \<lambda>env z. env \<in> list(A) & z = \<langle>env, not(rp`env and rq`env)\<rangle>)"
+         (M, \<lambda>env z. env \<in> list(A) \<and> z = \<langle>env, not(rp`env and rq`env)\<rangle>)"
 by (insert Nand_replacement, simp) 
 
 lemma (in M_satisfies) Forall_replacement':
    "\<lbrakk>M(A); M(rp)\<rbrakk>
     \<Longrightarrow> strong_replacement
         (M, \<lambda>env z.
-               env \<in> list(A) &
+               env \<in> list(A) \<and>
                z = \<langle>env, bool_of_o (\<forall>a\<in>A. rp ` Cons(a,env) = 1)\<rangle>)"
 by (insert Forall_replacement, simp) 
 
@@ -447,7 +447,7 @@
 
 lemma (in M_satisfies) fr_lam_replace:
    "\<lbrakk>M(g); M(A)\<rbrakk> \<Longrightarrow>
-    strong_replacement (M, \<lambda>x y. x \<in> formula &
+    strong_replacement (M, \<lambda>x y. x \<in> formula \<and>
             y = \<langle>x, 
                  formula_rec_case(satisfies_a(A),
                                   satisfies_b(A),
@@ -512,8 +512,8 @@
 (* is_depth_apply(M,h,p,z) \<equiv>
     \<exists>dp[M]. \<exists>sdp[M]. \<exists>hsdp[M]. 
       2        1         0
-        finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) &
-        fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z) *)
+        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(h,p,z) \<equiv>
@@ -555,7 +555,7 @@
              is_lambda(M, lA, 
                 \<lambda>env z. is_bool_of_o(M, 
                       \<exists>nx[M]. \<exists>ny[M]. 
-                       is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & nx \<in> ny, z),
+                       is_nth(M,x,env,nx) \<and> is_nth(M,y,env,ny) \<and> nx \<in> ny, z),
                 zz)  *)
 
 definition
@@ -607,7 +607,7 @@
     \<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, 
                 \<lambda>env z. is_bool_of_o(M, 
-                      \<exists>nx[M]. is_nth(M,x,env,nx) & is_nth(M,y,env,nx), z),
+                      \<exists>nx[M]. is_nth(M,x,env,nx) \<and> is_nth(M,y,env,nx), z),
                 zz) *)
 
 definition
@@ -655,9 +655,9 @@
 (* 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]. 
-                 (\<exists>rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) & 
-                 (\<exists>rq[M]. is_depth_apply(M,h,q,rq) & fun_apply(M,rq,env,hq)) & 
-                 (\<exists>pq[M]. is_and(M,hp,hq,pq) & is_not(M,pq,z)),
+                 (\<exists>rp[M]. is_depth_apply(M,h,p,rp) \<and> fun_apply(M,rp,env,hp)) \<and> 
+                 (\<exists>rq[M]. is_depth_apply(M,h,q,rq) \<and> fun_apply(M,rq,env,hq)) \<and> 
+                 (\<exists>pq[M]. is_and(M,hp,hq,pq) \<and> is_not(M,pq,z)),
                 zz) *)
 
 definition
@@ -706,7 +706,7 @@
 (* satisfies_is_d(M,A,h) \<equiv> 
     \<lambda>p zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, 
-                \<lambda>env z. \<exists>rp[M]. is_depth_apply(M,h,p,rp) & 
+                \<lambda>env z. \<exists>rp[M]. is_depth_apply(M,h,p,rp) \<and> 
                     is_bool_of_o(M, 
                            \<forall>x[M]. \<forall>xenv[M]. \<forall>hp[M]. 
                               x\<in>A \<longrightarrow> is_Cons(M,x,env,xenv) \<longrightarrow> 
@@ -836,8 +836,8 @@
     "\<lbrakk>L(A); x \<in> nat; y \<in> nat\<rbrakk>
      \<Longrightarrow> strong_replacement
          (L, \<lambda>env z. \<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. 
-              env \<in> list(A) & is_nth(L,x,env,nx) & is_nth(L,y,env,ny) & 
-              is_bool_of_o(L, nx \<in> ny, bo) &
+              env \<in> list(A) \<and> is_nth(L,x,env,nx) \<and> is_nth(L,y,env,ny) \<and> 
+              is_bool_of_o(L, nx \<in> ny, bo) \<and>
               pair(L, env, bo, z))"
 apply (rule strong_replacementI)
 apply (rule_tac u="{list(A),B,x,y}" 
@@ -866,8 +866,8 @@
     "\<lbrakk>L(A); x \<in> nat; y \<in> nat\<rbrakk>
      \<Longrightarrow> strong_replacement
          (L, \<lambda>env z. \<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. 
-              env \<in> list(A) & is_nth(L,x,env,nx) & is_nth(L,y,env,ny) & 
-              is_bool_of_o(L, nx = ny, bo) &
+              env \<in> list(A) \<and> is_nth(L,x,env,nx) \<and> is_nth(L,y,env,ny) \<and> 
+              is_bool_of_o(L, nx = ny, bo) \<and>
               pair(L, env, bo, z))"
 apply (rule strong_replacementI)
 apply (rule_tac u="{list(A),B,x,y}" 
@@ -898,9 +898,9 @@
     "\<lbrakk>L(A); L(rp); L(rq)\<rbrakk>
      \<Longrightarrow> strong_replacement
          (L, \<lambda>env z. \<exists>rpe[L]. \<exists>rqe[L]. \<exists>andpq[L]. \<exists>notpq[L]. 
-               fun_apply(L,rp,env,rpe) & fun_apply(L,rq,env,rqe) & 
-               is_and(L,rpe,rqe,andpq) & is_not(L,andpq,notpq) & 
-               env \<in> list(A) & pair(L, env, notpq, z))"
+               fun_apply(L,rp,env,rpe) \<and> fun_apply(L,rq,env,rqe) \<and> 
+               is_and(L,rpe,rqe,andpq) \<and> is_not(L,andpq,notpq) \<and> 
+               env \<in> list(A) \<and> pair(L, env, notpq, z))"
 apply (rule strong_replacementI)
 apply (rule_tac u="{list(A),B,rp,rq}" 
          in gen_separation_multi [OF Nand_Reflects],
@@ -933,12 +933,12 @@
    "\<lbrakk>L(A); L(rp)\<rbrakk>
     \<Longrightarrow> strong_replacement
         (L, \<lambda>env z. \<exists>bo[L]. 
-              env \<in> list(A) & 
+              env \<in> list(A) \<and> 
               is_bool_of_o (L, 
                             \<forall>a[L]. \<forall>co[L]. \<forall>rpco[L]. 
                                a\<in>A \<longrightarrow> is_Cons(L,a,env,co) \<longrightarrow>
                                fun_apply(L,rp,co,rpco) \<longrightarrow> number1(L, rpco), 
-                            bo) &
+                            bo) \<and>
               pair(L,env,bo,z))"
 apply (rule strong_replacementI)
 apply (rule_tac u="{A,list(A),B,rp}" 
@@ -974,20 +974,20 @@
 subsubsection\<open>The Lambda Replacement Case\<close>
 
 lemma formula_rec_lambda_replacement_Reflects:
- "REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B &
-     mem_formula(L,u) &
+ "REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B \<and>
+     mem_formula(L,u) \<and>
      (\<exists>c[L].
          is_formula_case
           (L, satisfies_is_a(L,A), satisfies_is_b(L,A),
            satisfies_is_c(L,A,g), satisfies_is_d(L,A,g),
-           u, c) &
+           u, c) \<and>
          pair(L,u,c,x)),
-  \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & mem_formula(##Lset(i),u) &
+  \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> mem_formula(##Lset(i),u) \<and>
      (\<exists>c \<in> Lset(i).
          is_formula_case
           (##Lset(i), satisfies_is_a(##Lset(i),A), satisfies_is_b(##Lset(i),A),
            satisfies_is_c(##Lset(i),A,g), satisfies_is_d(##Lset(i),A,g),
-           u, c) &
+           u, c) \<and>
          pair(##Lset(i),u,c,x))]"
 by (intro FOL_reflections function_reflections mem_formula_reflection
           is_formula_case_reflection satisfies_is_a_reflection
@@ -998,11 +998,11 @@
       \<comment> \<open>For the \<^term>\<open>transrec\<close>\<close>
    "\<lbrakk>L(g); L(A)\<rbrakk> \<Longrightarrow>
     strong_replacement (L, 
-       \<lambda>x y. mem_formula(L,x) &
+       \<lambda>x y. mem_formula(L,x) \<and>
              (\<exists>c[L]. is_formula_case(L, satisfies_is_a(L,A),
                                   satisfies_is_b(L,A),
                                   satisfies_is_c(L,A,g),
-                                  satisfies_is_d(L,A,g), x, c) &
+                                  satisfies_is_d(L,A,g), x, c) \<and>
              pair(L, x, c, y)))" 
 apply (rule strong_replacementI)
 apply (rule_tac u="{B,A,g}"
--- a/src/ZF/Constructible/Separation.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/Separation.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -18,12 +18,12 @@
 
 lemma Collect_conj_in_DPow:
      "\<lbrakk>{x\<in>A. P(x)} \<in> DPow(A);  {x\<in>A. Q(x)} \<in> DPow(A)\<rbrakk>
-      \<Longrightarrow> {x\<in>A. P(x) & Q(x)} \<in> DPow(A)"
+      \<Longrightarrow> {x\<in>A. P(x) \<and> Q(x)} \<in> DPow(A)"
 by (simp add: Int_in_DPow Collect_Int_Collect_eq [symmetric])
 
 lemma Collect_conj_in_DPow_Lset:
      "\<lbrakk>z \<in> Lset(j); {x \<in> Lset(j). P(x)} \<in> DPow(Lset(j))\<rbrakk>
-      \<Longrightarrow> {x \<in> Lset(j). x \<in> z & P(x)} \<in> DPow(Lset(j))"
+      \<Longrightarrow> {x \<in> Lset(j). x \<in> z \<and> P(x)} \<in> DPow(Lset(j))"
 apply (frule mem_Lset_imp_subset_Lset)
 apply (simp add: Collect_conj_in_DPow Collect_mem_eq
                  subset_Int_iff2 elem_subset_in_DPow)
@@ -43,7 +43,7 @@
           Ord(j);  z \<in> Lset(j)\<rbrakk> \<Longrightarrow> L({x \<in> z . P(x)})"
 apply (rule_tac i = "succ(j)" in L_I)
  prefer 2 apply simp
-apply (subgoal_tac "{x \<in> z. P(x)} = {x \<in> Lset(j). x \<in> z & (Q(x))}")
+apply (subgoal_tac "{x \<in> z. P(x)} = {x \<in> Lset(j). x \<in> z \<and> (Q(x))}")
  prefer 2
  apply (blast dest: mem_Lset_imp_subset_Lset)
 apply (simp add: Lset_succ Collect_conj_in_DPow_Lset)
@@ -118,14 +118,14 @@
 subsection\<open>Separation for Cartesian Product\<close>
 
 lemma cartprod_Reflects:
-     "REFLECTS[\<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>B & pair(L,x,y,z)),
-                \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A & (\<exists>y\<in>Lset(i). y\<in>B &
+     "REFLECTS[\<lambda>z. \<exists>x[L]. x\<in>A \<and> (\<exists>y[L]. y\<in>B \<and> pair(L,x,y,z)),
+                \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A \<and> (\<exists>y\<in>Lset(i). y\<in>B \<and>
                                    pair(##Lset(i),x,y,z))]"
 by (intro FOL_reflections function_reflections)
 
 lemma cartprod_separation:
      "\<lbrakk>L(A); L(B)\<rbrakk>
-      \<Longrightarrow> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>B & pair(L,x,y,z)))"
+      \<Longrightarrow> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A \<and> (\<exists>y[L]. y\<in>B \<and> pair(L,x,y,z)))"
 apply (rule gen_separation_multi [OF cartprod_Reflects, of "{A,B}"], auto)
 apply (rule_tac env="[A,B]" in DPow_LsetI)
 apply (rule sep_rules | simp)+
@@ -134,13 +134,13 @@
 subsection\<open>Separation for Image\<close>
 
 lemma image_Reflects:
-     "REFLECTS[\<lambda>y. \<exists>p[L]. p\<in>r & (\<exists>x[L]. x\<in>A & pair(L,x,y,p)),
-           \<lambda>i y. \<exists>p\<in>Lset(i). p\<in>r & (\<exists>x\<in>Lset(i). x\<in>A & pair(##Lset(i),x,y,p))]"
+     "REFLECTS[\<lambda>y. \<exists>p[L]. p\<in>r \<and> (\<exists>x[L]. x\<in>A \<and> pair(L,x,y,p)),
+           \<lambda>i y. \<exists>p\<in>Lset(i). p\<in>r \<and> (\<exists>x\<in>Lset(i). x\<in>A \<and> pair(##Lset(i),x,y,p))]"
 by (intro FOL_reflections function_reflections)
 
 lemma image_separation:
      "\<lbrakk>L(A); L(r)\<rbrakk>
-      \<Longrightarrow> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & (\<exists>x[L]. x\<in>A & pair(L,x,y,p)))"
+      \<Longrightarrow> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r \<and> (\<exists>x[L]. x\<in>A \<and> pair(L,x,y,p)))"
 apply (rule gen_separation_multi [OF image_Reflects, of "{A,r}"], auto)
 apply (rule_tac env="[A,r]" in DPow_LsetI)
 apply (rule sep_rules | simp)+
@@ -150,14 +150,14 @@
 subsection\<open>Separation for Converse\<close>
 
 lemma converse_Reflects:
-  "REFLECTS[\<lambda>z. \<exists>p[L]. p\<in>r & (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) & pair(L,y,x,z)),
-     \<lambda>i z. \<exists>p\<in>Lset(i). p\<in>r & (\<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i).
-                     pair(##Lset(i),x,y,p) & pair(##Lset(i),y,x,z))]"
+  "REFLECTS[\<lambda>z. \<exists>p[L]. p\<in>r \<and> (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) \<and> pair(L,y,x,z)),
+     \<lambda>i z. \<exists>p\<in>Lset(i). p\<in>r \<and> (\<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i).
+                     pair(##Lset(i),x,y,p) \<and> pair(##Lset(i),y,x,z))]"
 by (intro FOL_reflections function_reflections)
 
 lemma converse_separation:
      "L(r) \<Longrightarrow> separation(L,
-         \<lambda>z. \<exists>p[L]. p\<in>r & (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) & pair(L,y,x,z)))"
+         \<lambda>z. \<exists>p[L]. p\<in>r \<and> (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) \<and> pair(L,y,x,z)))"
 apply (rule gen_separation [OF converse_Reflects], simp)
 apply (rule_tac env="[r]" in DPow_LsetI)
 apply (rule sep_rules | simp)+
@@ -167,12 +167,12 @@
 subsection\<open>Separation for Restriction\<close>
 
 lemma restrict_Reflects:
-     "REFLECTS[\<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. pair(L,x,y,z)),
-        \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A & (\<exists>y\<in>Lset(i). pair(##Lset(i),x,y,z))]"
+     "REFLECTS[\<lambda>z. \<exists>x[L]. x\<in>A \<and> (\<exists>y[L]. pair(L,x,y,z)),
+        \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A \<and> (\<exists>y\<in>Lset(i). pair(##Lset(i),x,y,z))]"
 by (intro FOL_reflections function_reflections)
 
 lemma restrict_separation:
-   "L(A) \<Longrightarrow> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. pair(L,x,y,z)))"
+   "L(A) \<Longrightarrow> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A \<and> (\<exists>y[L]. pair(L,x,y,z)))"
 apply (rule gen_separation [OF restrict_Reflects], simp)
 apply (rule_tac env="[A]" in DPow_LsetI)
 apply (rule sep_rules | simp)+
@@ -183,18 +183,18 @@
 
 lemma comp_Reflects:
      "REFLECTS[\<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L].
-                  pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) &
-                  xy\<in>s & yz\<in>r,
+                  pair(L,x,z,xz) \<and> pair(L,x,y,xy) \<and> pair(L,y,z,yz) \<and>
+                  xy\<in>s \<and> yz\<in>r,
         \<lambda>i xz. \<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i). \<exists>z\<in>Lset(i). \<exists>xy\<in>Lset(i). \<exists>yz\<in>Lset(i).
-                  pair(##Lset(i),x,z,xz) & pair(##Lset(i),x,y,xy) &
-                  pair(##Lset(i),y,z,yz) & xy\<in>s & yz\<in>r]"
+                  pair(##Lset(i),x,z,xz) \<and> pair(##Lset(i),x,y,xy) \<and>
+                  pair(##Lset(i),y,z,yz) \<and> xy\<in>s \<and> yz\<in>r]"
 by (intro FOL_reflections function_reflections)
 
 lemma comp_separation:
      "\<lbrakk>L(r); L(s)\<rbrakk>
       \<Longrightarrow> separation(L, \<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L].
-                  pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) &
-                  xy\<in>s & yz\<in>r)"
+                  pair(L,x,z,xz) \<and> pair(L,x,y,xy) \<and> pair(L,y,z,yz) \<and>
+                  xy\<in>s \<and> yz\<in>r)"
 apply (rule gen_separation_multi [OF comp_Reflects, of "{r,s}"], auto)
 txt\<open>Subgoals after applying general ``separation'' rule:
      @{subgoals[display,indent=0,margin=65]}\<close>
@@ -208,12 +208,12 @@
 subsection\<open>Separation for Predecessors in an Order\<close>
 
 lemma pred_Reflects:
-     "REFLECTS[\<lambda>y. \<exists>p[L]. p\<in>r & pair(L,y,x,p),
-                    \<lambda>i y. \<exists>p \<in> Lset(i). p\<in>r & pair(##Lset(i),y,x,p)]"
+     "REFLECTS[\<lambda>y. \<exists>p[L]. p\<in>r \<and> pair(L,y,x,p),
+                    \<lambda>i y. \<exists>p \<in> Lset(i). p\<in>r \<and> pair(##Lset(i),y,x,p)]"
 by (intro FOL_reflections function_reflections)
 
 lemma pred_separation:
-     "\<lbrakk>L(r); L(x)\<rbrakk> \<Longrightarrow> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & pair(L,y,x,p))"
+     "\<lbrakk>L(r); L(x)\<rbrakk> \<Longrightarrow> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r \<and> pair(L,y,x,p))"
 apply (rule gen_separation_multi [OF pred_Reflects, of "{r,x}"], auto)
 apply (rule_tac env="[r,x]" in DPow_LsetI)
 apply (rule sep_rules | simp)+
@@ -223,12 +223,12 @@
 subsection\<open>Separation for the Membership Relation\<close>
 
 lemma Memrel_Reflects:
-     "REFLECTS[\<lambda>z. \<exists>x[L]. \<exists>y[L]. pair(L,x,y,z) & x \<in> y,
-            \<lambda>i z. \<exists>x \<in> Lset(i). \<exists>y \<in> Lset(i). pair(##Lset(i),x,y,z) & x \<in> y]"
+     "REFLECTS[\<lambda>z. \<exists>x[L]. \<exists>y[L]. pair(L,x,y,z) \<and> x \<in> y,
+            \<lambda>i z. \<exists>x \<in> Lset(i). \<exists>y \<in> Lset(i). pair(##Lset(i),x,y,z) \<and> x \<in> y]"
 by (intro FOL_reflections function_reflections)
 
 lemma Memrel_separation:
-     "separation(L, \<lambda>z. \<exists>x[L]. \<exists>y[L]. pair(L,x,y,z) & x \<in> y)"
+     "separation(L, \<lambda>z. \<exists>x[L]. \<exists>y[L]. pair(L,x,y,z) \<and> x \<in> y)"
 apply (rule gen_separation [OF Memrel_Reflects nonempty])
 apply (rule_tac env="[]" in DPow_LsetI)
 apply (rule sep_rules | simp)+
@@ -238,19 +238,19 @@
 subsection\<open>Replacement for FunSpace\<close>
 
 lemma funspace_succ_Reflects:
- "REFLECTS[\<lambda>z. \<exists>p[L]. p\<in>A & (\<exists>f[L]. \<exists>b[L]. \<exists>nb[L]. \<exists>cnbf[L].
-            pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) &
+ "REFLECTS[\<lambda>z. \<exists>p[L]. p\<in>A \<and> (\<exists>f[L]. \<exists>b[L]. \<exists>nb[L]. \<exists>cnbf[L].
+            pair(L,f,b,p) \<and> pair(L,n,b,nb) \<and> is_cons(L,nb,f,cnbf) \<and>
             upair(L,cnbf,cnbf,z)),
-        \<lambda>i z. \<exists>p \<in> Lset(i). p\<in>A & (\<exists>f \<in> Lset(i). \<exists>b \<in> Lset(i).
+        \<lambda>i z. \<exists>p \<in> Lset(i). p\<in>A \<and> (\<exists>f \<in> Lset(i). \<exists>b \<in> Lset(i).
               \<exists>nb \<in> Lset(i). \<exists>cnbf \<in> Lset(i).
-                pair(##Lset(i),f,b,p) & pair(##Lset(i),n,b,nb) &
-                is_cons(##Lset(i),nb,f,cnbf) & upair(##Lset(i),cnbf,cnbf,z))]"
+                pair(##Lset(i),f,b,p) \<and> pair(##Lset(i),n,b,nb) \<and>
+                is_cons(##Lset(i),nb,f,cnbf) \<and> upair(##Lset(i),cnbf,cnbf,z))]"
 by (intro FOL_reflections function_reflections)
 
 lemma funspace_succ_replacement:
      "L(n) \<Longrightarrow>
       strong_replacement(L, \<lambda>p z. \<exists>f[L]. \<exists>b[L]. \<exists>nb[L]. \<exists>cnbf[L].
-                pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) &
+                pair(L,f,b,p) \<and> pair(L,n,b,nb) \<and> is_cons(L,nb,f,cnbf) \<and>
                 upair(L,cnbf,cnbf,z))"
 apply (rule strong_replacementI)
 apply (rule_tac u="{n,B}" in gen_separation_multi [OF funspace_succ_Reflects], 
@@ -264,13 +264,13 @@
 
 lemma is_recfun_reflects:
   "REFLECTS[\<lambda>x. \<exists>xa[L]. \<exists>xb[L].
-                pair(L,x,a,xa) & xa \<in> r & pair(L,x,b,xb) & xb \<in> r &
-                (\<exists>fx[L]. \<exists>gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) &
+                pair(L,x,a,xa) \<and> xa \<in> r \<and> pair(L,x,b,xb) \<and> xb \<in> r \<and>
+                (\<exists>fx[L]. \<exists>gx[L]. fun_apply(L,f,x,fx) \<and> fun_apply(L,g,x,gx) \<and>
                                    fx \<noteq> gx),
    \<lambda>i x. \<exists>xa \<in> Lset(i). \<exists>xb \<in> Lset(i).
-          pair(##Lset(i),x,a,xa) & xa \<in> r & pair(##Lset(i),x,b,xb) & xb \<in> r &
-                (\<exists>fx \<in> Lset(i). \<exists>gx \<in> Lset(i). fun_apply(##Lset(i),f,x,fx) &
-                  fun_apply(##Lset(i),g,x,gx) & fx \<noteq> gx)]"
+          pair(##Lset(i),x,a,xa) \<and> xa \<in> r \<and> pair(##Lset(i),x,b,xb) \<and> xb \<in> r \<and>
+                (\<exists>fx \<in> Lset(i). \<exists>gx \<in> Lset(i). fun_apply(##Lset(i),f,x,fx) \<and>
+                  fun_apply(##Lset(i),g,x,gx) \<and> fx \<noteq> gx)]"
 by (intro FOL_reflections function_reflections fun_plus_reflections)
 
 lemma is_recfun_separation:
@@ -278,8 +278,8 @@
      "\<lbrakk>L(r); L(f); L(g); L(a); L(b)\<rbrakk>
      \<Longrightarrow> separation(L,
             \<lambda>x. \<exists>xa[L]. \<exists>xb[L].
-                pair(L,x,a,xa) & xa \<in> r & pair(L,x,b,xb) & xb \<in> r &
-                (\<exists>fx[L]. \<exists>gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) &
+                pair(L,x,a,xa) \<and> xa \<in> r \<and> pair(L,x,b,xb) \<and> xb \<in> r \<and>
+                (\<exists>fx[L]. \<exists>gx[L]. fun_apply(L,f,x,fx) \<and> fun_apply(L,g,x,gx) \<and>
                                    fx \<noteq> gx))"
 apply (rule gen_separation_multi [OF is_recfun_reflects, of "{r,f,g,a,b}"], 
             auto)
--- a/src/ZF/Constructible/WF_absolute.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/WF_absolute.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -12,7 +12,7 @@
   rtrancl_alt :: "[i,i]=>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> &  f`0 = x & f`n = y) &
+                 (\<exists>x y. p = <x,y> \<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]:
@@ -63,14 +63,14 @@
     \<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]. 
-               omega(M,nnat) & n\<in>nnat & successor(M,n,n') &
-               (\<exists>f[M]. typed_function(M,n',A,f) &
-                (\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) &
-                  fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) &
+               omega(M,nnat) \<and> n\<in>nnat \<and> successor(M,n,n') \<and>
+               (\<exists>f[M]. typed_function(M,n',A,f) \<and>
+                (\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) \<and> empty(M,zero) \<and>
+                  fun_apply(M,f,zero,x) \<and> fun_apply(M,f,n,y)) \<and>
                   (\<forall>j[M]. j\<in>n \<longrightarrow> 
                     (\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M]. 
-                      fun_apply(M,f,j,fj) & successor(M,j,sj) &
-                      fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"
+                      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 :: "[i=>o,i,i] => o" where
@@ -81,7 +81,7 @@
 definition
   tran_closure :: "[i=>o,i,i] => o" where
     "tran_closure(M,r,t) \<equiv>
-         \<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)"
+         \<exists>s[M]. rtran_closure(M,r,s) \<and> composition(M,r,s,t)"
     
 locale M_trancl = M_basic +
   assumes rtrancl_separation:
@@ -90,15 +90,15 @@
          "\<lbrakk>M(r); M(Z)\<rbrakk> \<Longrightarrow> 
           separation (M, \<lambda>x. 
               \<exists>w[M]. \<exists>wx[M]. \<exists>rp[M]. 
-               w \<in> Z & pair(M,w,x,wx) & tran_closure(M,r,rp) & wx \<in> rp)"
+               w \<in> Z \<and> pair(M,w,x,wx) \<and> tran_closure(M,r,rp) \<and> wx \<in> rp)"
       and M_nat [iff] : "M(nat)"
 
 lemma (in M_trancl) rtran_closure_mem_iff:
      "\<lbrakk>M(A); M(r); M(p)\<rbrakk>
       \<Longrightarrow> rtran_closure_mem(M,A,r,p) \<longleftrightarrow>
-          (\<exists>n[M]. n\<in>nat & 
-           (\<exists>f[M]. f \<in> succ(n) -> A &
-            (\<exists>x[M]. \<exists>y[M]. p = <x,y> & f`0 = x & f`n = y) &
+          (\<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>
                            (\<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 & <w,x> \<in> r^+)"
+     "\<lbrakk>M(r); M(Z)\<rbrakk> \<Longrightarrow> separation (M, \<lambda>x. \<exists>w[M]. w \<in> Z \<and> <w,x> \<in> r^+)"
 by (insert wellfounded_trancl_separation [of r Z], simp) 
 
 text\<open>Alternative proof of \<open>wf_on_trancl\<close>; inspiration for the
@@ -145,7 +145,7 @@
 lemma "\<lbrakk>wf[A](r);  r-``A \<subseteq> A\<rbrakk> \<Longrightarrow> wf[A](r^+)"
 apply (simp add: wf_on_def wf_def)
 apply (safe)
-apply (drule_tac x = "{x\<in>A. \<exists>w. \<langle>w,x\<rangle> \<in> r^+ & w \<in> Z}" in spec)
+apply (drule_tac x = "{x\<in>A. \<exists>w. \<langle>w,x\<rangle> \<in> r^+ \<and> w \<in> Z}" in spec)
 apply (blast elim: tranclE)
 done
 
@@ -155,10 +155,10 @@
 apply (simp add: wellfounded_on_def)
 apply (safe intro!: equalityI)
 apply (rename_tac Z x)
-apply (subgoal_tac "M({x\<in>A. \<exists>w[M]. w \<in> Z & \<langle>w,x\<rangle> \<in> r^+})")
+apply (subgoal_tac "M({x\<in>A. \<exists>w[M]. w \<in> Z \<and> \<langle>w,x\<rangle> \<in> r^+})")
  prefer 2
  apply (blast intro: wellfounded_trancl_separation') 
-apply (drule_tac x = "{x\<in>A. \<exists>w[M]. w \<in> Z & \<langle>w,x\<rangle> \<in> r^+}" in rspec, safe)
+apply (drule_tac x = "{x\<in>A. \<exists>w[M]. w \<in> Z \<and> \<langle>w,x\<rangle> \<in> r^+}" in rspec, safe)
 apply (blast dest: transM, simp)
 apply (rename_tac y w)
 apply (drule_tac x=w in bspec, assumption, clarify)
@@ -183,12 +183,12 @@
 lemma (in M_trancl) wfrec_relativize:
   "\<lbrakk>wf(r); M(a); M(r);  
      strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
-          pair(M,x,y,z) & 
-          is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
+          pair(M,x,y,z) \<and> 
+          is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) \<and> 
           y = H(x, restrict(g, r -`` {x}))); 
      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
    \<Longrightarrow> wfrec(r,a,H) = z \<longleftrightarrow> 
-       (\<exists>f[M]. is_recfun(r^+, a, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) & 
+       (\<exists>f[M]. is_recfun(r^+, a, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) \<and> 
             z = H(a,restrict(f,r-``{a})))"
 apply (frule wf_trancl) 
 apply (simp add: wftrec_def wfrec_def, safe)
@@ -207,7 +207,7 @@
   "\<lbrakk>wf(r);  trans(r);  relation(r);  M(r);  M(a);
      wfrec_replacement(M,MH,r);  relation2(M,MH,H);
      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
-   \<Longrightarrow> wfrec(r,a,H) = z \<longleftrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f) & z = H(a,f))" 
+   \<Longrightarrow> wfrec(r,a,H) = z \<longleftrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f) \<and> z = H(a,f))" 
 apply (frule wfrec_replacement', assumption+) 
 apply (simp cong: is_recfun_cong
            add: wfrec_relativize trancl_eq_r
@@ -227,7 +227,7 @@
      wfrec_replacement(M,MH,r);  relation2(M,MH,H);
      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
    \<Longrightarrow> y = <x, wfrec(r, x, H)> \<longleftrightarrow> 
-       (\<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
+       (\<exists>f[M]. is_recfun(r,x,H,f) \<and> y = <x, H(x,f)>)"
 apply safe 
  apply (simp add: trans_wfrec_relativize [THEN iff_sym, of concl: _ x]) 
 txt\<open>converse direction\<close>
@@ -255,9 +255,9 @@
 text\<open>Eliminates one instance of replacement.\<close>
 lemma (in M_trancl) wfrec_replacement_iff:
      "strong_replacement(M, \<lambda>x z. 
-          \<exists>y[M]. pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g))) \<longleftrightarrow>
+          \<exists>y[M]. pair(M,x,y,z) \<and> (\<exists>g[M]. is_recfun(r,x,H,g) \<and> y = H(x,g))) \<longleftrightarrow>
       strong_replacement(M, 
-           \<lambda>x y. \<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
+           \<lambda>x y. \<exists>f[M]. is_recfun(r,x,H,f) \<and> y = <x, H(x,f)>)"
 apply simp 
 apply (rule strong_replacement_cong, blast) 
 done
@@ -278,12 +278,12 @@
 lemma (in M_trancl) eq_pair_wfrec_iff:
   "\<lbrakk>wf(r);  M(r);  M(y); 
      strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
-          pair(M,x,y,z) & 
-          is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
+          pair(M,x,y,z) \<and> 
+          is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) \<and> 
           y = H(x, restrict(g, r -`` {x}))); 
      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
    \<Longrightarrow> y = <x, wfrec(r, x, H)> \<longleftrightarrow> 
-       (\<exists>f[M]. is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) & 
+       (\<exists>f[M]. is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) \<and> 
             y = <x, H(x,restrict(f,r-``{x}))>)"
 apply safe  
  apply (simp add: wfrec_relativize [THEN iff_sym, of concl: _ x]) 
--- a/src/ZF/Constructible/WFrec.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/WFrec.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -21,7 +21,7 @@
 text\<open>Expresses \<open>is_recfun\<close> as a recursion equation\<close>
 lemma is_recfun_iff_equation:
      "is_recfun(r,a,H,f) \<longleftrightarrow>
-           f \<in> r -`` {a} \<rightarrow> range(f) &
+           f \<in> r -`` {a} \<rightarrow> range(f) \<and>
            (\<forall>x \<in> r-``{a}. f`x = H(x, restrict(f, r-``{x})))"  
 apply (rule iffI) 
  apply (simp add: is_recfun_type apply_recfun Ball_def vimage_singleton_iff, 
@@ -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 & z = <x, H(x, restrict(f, r-``{x}))>))"
+        (\<exists>x[M]. <x,a> \<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,8 +174,8 @@
        \<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 &
-            (\<exists>y[M]. b = \<langle>x,y\<rangle> & (\<exists>g[M]. is_recfun(r,x,H,g) \<and> y = H(x,g))));
+           (\<exists>x[M]. <x,a1> \<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") 
@@ -201,7 +201,7 @@
 lemma (in M_basic) univalent_is_recfun:
      "\<lbrakk>wellfounded(M,r); trans(r); M(r)\<rbrakk>
       \<Longrightarrow> univalent (M, A, \<lambda>x p. 
-              \<exists>y[M]. p = \<langle>x,y\<rangle> & (\<exists>f[M]. is_recfun(r,x,H,f) & y = H(x,f)))"
+              \<exists>y[M]. p = \<langle>x,y\<rangle> \<and> (\<exists>f[M]. is_recfun(r,x,H,f) \<and> y = H(x,f)))"
 apply (simp add: univalent_def) 
 apply (blast dest: is_recfun_functional) 
 done
@@ -213,7 +213,7 @@
     "\<lbrakk>\<forall>y. \<langle>y, a1\<rangle> \<in> r \<longrightarrow> (\<exists>f[M]. is_recfun(r, y, H, f)); 
        wellfounded(M,r); trans(r); M(r); M(a1);
        strong_replacement(M, \<lambda>x z. 
-              \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); 
+              \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) \<and> is_recfun(r,x,H,g) \<and> y = H(x,g)); 
        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>   
       \<Longrightarrow> \<exists>f[M]. is_recfun(r,a1,H,f)"
 apply (drule_tac A="r-``{a1}" in strong_replacementD)
@@ -246,7 +246,7 @@
     "\<lbrakk>wellfounded(M,r);  trans(r);  
        separation(M, \<lambda>x. \<not> (\<exists>f[M]. is_recfun(r, x, H, f)));
        strong_replacement(M, \<lambda>x z. 
-          \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); 
+          \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) \<and> is_recfun(r,x,H,g) \<and> y = H(x,g)); 
        M(r);  M(a);  
        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>   
       \<Longrightarrow> \<exists>f[M]. is_recfun(r,a,H,f)"
@@ -257,7 +257,7 @@
 lemma (in M_basic) wf_exists_is_recfun [rule_format]:
     "\<lbrakk>wf(r);  trans(r);  M(r);  
        strong_replacement(M, \<lambda>x z. 
-         \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); 
+         \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) \<and> is_recfun(r,x,H,g) \<and> y = H(x,g)); 
        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>   
       \<Longrightarrow> M(a) \<longrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f))"
 apply (rule wf_induct, assumption+)
@@ -275,20 +275,20 @@
    "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]. 
-               pair(M,x,y,z) & pair(M,x,a,xa) & upair(M,x,x,sx) &
-               pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) &
-               xa \<in> r & MH(x, f_r_sx, y))"
+               pair(M,x,y,z) \<and> pair(M,x,a,xa) \<and> upair(M,x,x,sx) \<and>
+               pre_image(M,r,sx,r_sx) \<and> restriction(M,f,r_sx,f_r_sx) \<and>
+               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(M,MH,r,a,z) \<equiv> 
-      \<exists>f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,z)"
+      \<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(M,MH,r) \<equiv>
         strong_replacement(M, 
-             \<lambda>x z. \<exists>y[M]. pair(M,x,y,z) & is_wfrec(M,MH,r,x,y))"
+             \<lambda>x z. \<exists>y[M]. pair(M,x,y,z) \<and> is_wfrec(M,MH,r,x,y))"
 
 lemma (in M_basic) is_recfun_abs:
      "\<lbrakk>\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g));  M(r); M(a); M(f); 
@@ -309,7 +309,7 @@
      "\<lbrakk>\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)); 
          relation2(M,MH,H);  M(r); M(a); M(z)\<rbrakk>
       \<Longrightarrow> is_wfrec(M,MH,r,a,z) \<longleftrightarrow> 
-          (\<exists>g[M]. is_recfun(r,a,H,g) & z = H(a,g))"
+          (\<exists>g[M]. is_recfun(r,a,H,g) \<and> z = H(a,g))"
 by (simp add: is_wfrec_def relation2_def is_recfun_abs)
 
 text\<open>Relating \<^term>\<open>wfrec_replacement\<close> to native constructs\<close>
@@ -318,7 +318,7 @@
      \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)); 
      relation2(M,MH,H);  M(r)\<rbrakk> 
    \<Longrightarrow> strong_replacement(M, \<lambda>x z. \<exists>y[M]. 
-                pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g)))"
+                pair(M,x,y,z) \<and> (\<exists>g[M]. is_recfun(r,x,H,g) \<and> y = H(x,g)))"
 by (simp add: wfrec_replacement_def is_wfrec_abs) 
 
 lemma wfrec_replacement_cong [cong]:
--- a/src/ZF/Constructible/Wellorderings.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Constructible/Wellorderings.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -34,18 +34,18 @@
   wellfounded :: "[i=>o,i]=>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 & \<not>(\<exists>z[M]. z\<in>x & <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> <z,y> \<in> r))"
 definition
   wellfounded_on :: "[i=>o,i,i]=>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 & \<not>(\<exists>z[M]. z\<in>x & <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> <z,y> \<in> r))"
 
 definition
   wellordered :: "[i=>o,i,i]=>o" where
     \<comment> \<open>linear and wellfounded on \<open>A\<close>\<close>
     "wellordered(M,A,r) \<equiv> 
-        transitive_rel(M,A,r) & linear_rel(M,A,r) & wellfounded_on(M,A,r)"
+        transitive_rel(M,A,r) \<and> linear_rel(M,A,r) \<and> wellfounded_on(M,A,r)"
 
 
 subsubsection \<open>Trivial absoluteness proofs\<close>
@@ -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) & (\<forall>y. <y,x> \<in> r \<longrightarrow> P(y)) \<longrightarrow> P(x)\<rbrakk>
+         \<forall>x. M(x) \<and> (\<forall>y. <y,x> \<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) & (\<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. <y,x> \<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)
@@ -188,7 +188,7 @@
 done
 
 lemma (in M_basic) M_Memrel_iff:
-     "M(A) \<Longrightarrow> Memrel(A) = {z \<in> A*A. \<exists>x[M]. \<exists>y[M]. z = \<langle>x,y\<rangle> & x \<in> y}"
+     "M(A) \<Longrightarrow> Memrel(A) = {z \<in> A*A. \<exists>x[M]. \<exists>y[M]. z = \<langle>x,y\<rangle> \<and> x \<in> y}"
 unfolding Memrel_def by (blast dest: transM)
 
 lemma (in M_basic) Memrel_closed [intro,simp]: 
--- a/src/ZF/Epsilon.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Epsilon.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -359,8 +359,8 @@
 
 lemma def_transrec2:
      "(\<And>x. f(x)\<equiv>transrec2(x,a,b))
-      \<Longrightarrow> f(0) = a &
-          f(succ(i)) = b(i, f(i)) &
+      \<Longrightarrow> f(0) = a \<and>
+          f(succ(i)) = b(i, f(i)) \<and>
           (Limit(K) \<longrightarrow> f(K) = (\<Union>j<K. f(j)))"
 by (simp add: transrec2_Limit)
 
--- a/src/ZF/EquivClass.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/EquivClass.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -92,7 +92,7 @@
 by (unfold equiv_def, blast)
 
 lemma equiv_class_eq_iff:
-     "equiv(A,r) \<Longrightarrow> <x,y>: r \<longleftrightarrow> r``{x} = r``{y} & x \<in> A & y \<in> A"
+     "equiv(A,r) \<Longrightarrow> <x,y>: 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:
--- a/src/ZF/Finite.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Finite.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -183,7 +183,7 @@
 
 lemma FiniteFun_Collect_iff:
      "f \<in> FiniteFun(A, {y \<in> B. P(y)})
-      \<longleftrightarrow> f \<in> FiniteFun(A,B) & (\<forall>x\<in>domain(f). P(f`x))"
+      \<longleftrightarrow> f \<in> FiniteFun(A,B) \<and> (\<forall>x\<in>domain(f). P(f`x))"
 apply auto
 apply (blast intro: FiniteFun_mono [THEN [2] rev_subsetD])
 apply (blast dest: Pair_mem_PiD FiniteFun_is_fun)
--- a/src/ZF/Fixedpt.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Fixedpt.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -10,7 +10,7 @@
 definition 
   (*monotone operator from Pow(D) to itself*)
   bnd_mono :: "[i,i=>i]=>o"  where
-     "bnd_mono(D,h) \<equiv> h(D)<=D & (\<forall>W X. W<=X \<longrightarrow> X<=D \<longrightarrow> h(W) \<subseteq> h(X))"
+     "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
--- a/src/ZF/IMP/Denotation.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/IMP/Denotation.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -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 & y \<in> loc->nat"
+    "\<lbrakk><x,y> \<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/Induct/Binary_Trees.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Induct/Binary_Trees.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -20,7 +20,7 @@
 lemma Br_neq_left: "l \<in> bt(A) \<Longrightarrow> Br(x, l, r) \<noteq> l"
   by (induct arbitrary: x r set: bt) auto
 
-lemma Br_iff: "Br(a, l, r) = Br(a', l', r') \<longleftrightarrow> a = a' & l = l' & r = r'"
+lemma Br_iff: "Br(a, l, r) = Br(a', l', r') \<longleftrightarrow> a = a' \<and> l = l' \<and> r = r'"
   \<comment> \<open>Proving a freeness theorem.\<close>
   by (fast elim!: bt.free_elims)
 
--- a/src/ZF/Induct/Comb.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Induct/Comb.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -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 & <y',z> \<in> r))"
+    \<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))"
 
 
 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^+ & <y,z>: r)"
+    \<forall>y'. <x,y'>:r \<longrightarrow> (\<exists>z. <y',z>: r^+ \<and> <y,z>: r)"
   apply (unfold diamond_def)
   apply (erule trancl_induct)
    apply (blast intro: r_into_trancl)
@@ -153,7 +153,7 @@
 lemma I_contract_E: "I \<rightarrow>\<^sup>1 r \<Longrightarrow> P"
   by (auto simp add: I_def)
 
-lemma K1_contractD: "K\<bullet>p \<rightarrow>\<^sup>1 r \<Longrightarrow> (\<exists>q. r = K\<bullet>q & p \<rightarrow>\<^sup>1 q)"
+lemma K1_contractD: "K\<bullet>p \<rightarrow>\<^sup>1 r \<Longrightarrow> (\<exists>q. r = K\<bullet>q \<and> p \<rightarrow>\<^sup>1 q)"
   by auto
 
 lemma Ap_reduce1: "\<lbrakk>p \<rightarrow> q;  r \<in> comb\<rbrakk> \<Longrightarrow> p\<bullet>r \<rightarrow> q\<bullet>r"
@@ -215,15 +215,15 @@
 subsection \<open>Basic properties of parallel contraction\<close>
 
 lemma K1_parcontractD [dest!]:
-    "K\<bullet>p \<Rrightarrow>\<^sup>1 r \<Longrightarrow> (\<exists>p'. r = K\<bullet>p' & p \<Rrightarrow>\<^sup>1 p')"
+    "K\<bullet>p \<Rrightarrow>\<^sup>1 r \<Longrightarrow> (\<exists>p'. r = K\<bullet>p' \<and> p \<Rrightarrow>\<^sup>1 p')"
   by auto
 
 lemma S1_parcontractD [dest!]:
-    "S\<bullet>p \<Rrightarrow>\<^sup>1 r \<Longrightarrow> (\<exists>p'. r = S\<bullet>p' & p \<Rrightarrow>\<^sup>1 p')"
+    "S\<bullet>p \<Rrightarrow>\<^sup>1 r \<Longrightarrow> (\<exists>p'. r = S\<bullet>p' \<and> p \<Rrightarrow>\<^sup>1 p')"
   by auto
 
 lemma S2_parcontractD [dest!]:
-    "S\<bullet>p\<bullet>q \<Rrightarrow>\<^sup>1 r \<Longrightarrow> (\<exists>p' q'. r = S\<bullet>p'\<bullet>q' & p \<Rrightarrow>\<^sup>1 p' & q \<Rrightarrow>\<^sup>1 q')"
+    "S\<bullet>p\<bullet>q \<Rrightarrow>\<^sup>1 r \<Longrightarrow> (\<exists>p' q'. r = S\<bullet>p'\<bullet>q' \<and> p \<Rrightarrow>\<^sup>1 p' \<and> q \<Rrightarrow>\<^sup>1 q')"
   by auto
 
 lemma diamond_parcontract: "diamond(parcontract)"
--- a/src/ZF/Induct/FoldSet.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Induct/FoldSet.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -38,7 +38,7 @@
 by (auto elim: equalityE)
 
 lemma cons_lemma2: "\<lbrakk>cons(x, B)=cons(y, C); x\<noteq>y; x\<notin>B; y\<notin>C\<rbrakk>  
-    \<Longrightarrow>  B - {y} = C-{x} & x\<in>C & y\<in>B"
+    \<Longrightarrow>  B - {y} = C-{x} \<and> x\<in>C \<and> y\<in>B"
 apply (auto elim: equalityE)
 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) & C<=A"
+     "<C, x>\<in>fold_set(A, B, f, e) \<Longrightarrow> <C, x>\<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
@@ -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) & v = f(c, y))"
+         (\<exists>y. <C, y> \<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+)
@@ -296,7 +296,7 @@
  prefer 2 apply blast
 apply (subgoal_tac "C (x) \<inter> (\<Union>i\<in>B. C (i)) = 0")
  prefer 2 apply blast
-apply (subgoal_tac "Finite (\<Union>i\<in>B. C (i)) & Finite (C (x)) & Finite (B) ")
+apply (subgoal_tac "Finite (\<Union>i\<in>B. C (i)) \<and> Finite (C (x)) \<and> Finite (B) ")
 apply (simp (no_asm_simp) add: setsum_Un_disjoint)
 apply (auto intro: Finite_Union Finite_RepFun)
 done
--- a/src/ZF/Induct/ListN.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Induct/ListN.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -24,7 +24,7 @@
 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) & length(l)=n"
+lemma listn_iff: "<n,l> \<in> listn(A) \<longleftrightarrow> l \<in> list(A) \<and> length(l)=n"
   apply (rule iffI)
    apply (erule listn.induct)
     apply auto
--- a/src/ZF/Induct/Multiset.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Induct/Multiset.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -26,7 +26,7 @@
 definition
   (* M is a multiset *)
   multiset :: "i => o"  where
-  "multiset(M) \<equiv> \<exists>A. M \<in> A -> nat-{0} & Finite(A)"
+  "multiset(M) \<equiv> \<exists>A. M \<in> A -> nat-{0} \<and> Finite(A)"
 
 definition
   mset_of :: "i=>i"  where
@@ -42,7 +42,7 @@
   (*convert a function to a multiset by eliminating 0*)
   normalize :: "i => i"  where
   "normalize(f) \<equiv>
-       if (\<exists>A. f \<in> A -> nat & Finite(A)) then
+       if (\<exists>A. f \<in> A -> nat \<and> Finite(A)) then
             funrestrict(f, {x \<in> mset_of(f). 0 < f`x})
        else 0"
 
@@ -87,7 +87,7 @@
   "multirel1(A, r) \<equiv>
      {<M, N> \<in> Mult(A)*Mult(A).
       \<exists>a \<in> A. \<exists>M0 \<in> Mult(A). \<exists>K \<in> Mult(A).
-      N=M0 +# {#a#} & M=M0 +# K & (\<forall>b \<in> mset_of(K). <b,a> \<in> r)}"
+      N=M0 +# {#a#} \<and> M=M0 +# K \<and> (\<forall>b \<in> mset_of(K). <b,a> \<in> r)}"
 
 definition
   multirel :: "[i, i] => i"  where
@@ -97,15 +97,15 @@
 
 definition
   omultiset :: "i => o"  where
-  "omultiset(M) \<equiv> \<exists>i. Ord(i) & M \<in> Mult(field(Memrel(i)))"
+  "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) & <M, N> \<in> multirel(field(Memrel(i)), Memrel(i))"
+  "M <# N \<equiv>  \<exists>i. Ord(i) \<and> <M, N> \<in> multirel(field(Memrel(i)), Memrel(i))"
 
 definition
   mle  :: "[i, i] => o"  (infixl \<open><#=\<close> 50)  where
-  "M <#= N \<equiv> (omultiset(M) & M = N) | M <# N"
+  "M <#= N \<equiv> (omultiset(M) \<and> M = N) | M <# N"
 
 
 subsection\<open>Properties of the original "restrict" from ZF.thy\<close>
@@ -142,7 +142,7 @@
 
 text\<open>A useful simplification rule\<close>
 lemma multiset_fun_iff:
-     "(f \<in> A -> nat-{0}) \<longleftrightarrow> f \<in> A->nat&(\<forall>a \<in> A. f`a \<in> nat & 0 < f`a)"
+     "(f \<in> A -> nat-{0}) \<longleftrightarrow> f \<in> A->nat\<and>(\<forall>a \<in> A. f`a \<in> nat \<and> 0 < f`a)"
 apply safe
 apply (rule_tac [4] B1 = "range (f) " in Pi_mono [THEN subsetD])
 apply (auto intro!: Ord_0_lt
@@ -159,7 +159,7 @@
 apply (simp_all (no_asm_simp) add: multiset_fun_iff)
 done
 
-lemma Mult_into_multiset: "M \<in> Mult(A) \<Longrightarrow> multiset(M) & mset_of(M)\<subseteq>A"
+lemma Mult_into_multiset: "M \<in> Mult(A) \<Longrightarrow> multiset(M) \<and> mset_of(M)\<subseteq>A"
 apply (simp add: multiset_def mset_of_def)
 apply (frule FiniteFun_is_fun)
 apply (drule FiniteFun_domain_Fin)
@@ -168,7 +168,7 @@
 apply (blast intro: Fin_into_Finite)
 done
 
-lemma Mult_iff_multiset: "M \<in> Mult(A) \<longleftrightarrow> multiset(M) & mset_of(M)\<subseteq>A"
+lemma Mult_iff_multiset: "M \<in> Mult(A) \<longleftrightarrow> multiset(M) \<and> mset_of(M)\<subseteq>A"
 by (blast dest: Mult_into_multiset intro: multiset_into_Mult)
 
 lemma multiset_iff_Mult_mset_of: "multiset(M) \<longleftrightarrow> M \<in> Mult(mset_of(M))"
@@ -205,7 +205,7 @@
 
 (* msingle *)
 
-lemma msingle_not_0 [iff]: "{#a#} \<noteq> 0 & 0 \<noteq> {#a#}"
+lemma msingle_not_0 [iff]: "{#a#} \<noteq> 0 \<and> 0 \<noteq> {#a#}"
 by (simp add: msingle_def)
 
 lemma msingle_eq_iff [iff]: "({#a#} = {#b#}) \<longleftrightarrow>  (a = b)"
@@ -223,7 +223,7 @@
 
 lemma normalize_idem [simp]: "normalize(normalize(f)) = normalize(f)"
 apply (simp add: normalize_def funrestrict_def mset_of_def)
-apply (case_tac "\<exists>A. f \<in> A -> nat & Finite (A) ")
+apply (case_tac "\<exists>A. f \<in> A -> nat \<and> Finite (A) ")
 apply clarify
 apply (drule_tac x = "{x \<in> domain (f) . 0 < f ` x}" in spec)
 apply auto
@@ -259,7 +259,7 @@
 
 (* Union *)
 
-lemma munion_0 [simp]: "multiset(M) \<Longrightarrow> M +# 0 = M & 0 +# M = M"
+lemma munion_0 [simp]: "multiset(M) \<Longrightarrow> M +# 0 = M \<and> 0 +# M = M"
 apply (simp add: multiset_def)
 apply (auto simp add: munion_def mset_of_def)
 done
@@ -424,10 +424,10 @@
 
 (** More algebraic properties of multisets **)
 
-lemma munion_eq_0_iff [simp]: "\<lbrakk>multiset(M); multiset(N)\<rbrakk>\<Longrightarrow>(M +# N =0) \<longleftrightarrow> (M=0 & N=0)"
+lemma munion_eq_0_iff [simp]: "\<lbrakk>multiset(M); multiset(N)\<rbrakk>\<Longrightarrow>(M +# N =0) \<longleftrightarrow> (M=0 \<and> N=0)"
 by (auto simp add: multiset_equality)
 
-lemma empty_eq_munion_iff [simp]: "\<lbrakk>multiset(M); multiset(N)\<rbrakk>\<Longrightarrow>(0=M +# N) \<longleftrightarrow> (M=0 & N=0)"
+lemma empty_eq_munion_iff [simp]: "\<lbrakk>multiset(M); multiset(N)\<rbrakk>\<Longrightarrow>(0=M +# N) \<longleftrightarrow> (M=0 \<and> N=0)"
 apply (rule iffI, drule sym)
 apply (simp_all add: multiset_equality)
 done
@@ -440,12 +440,12 @@
   "\<lbrakk>multiset(K); multiset(M); multiset(N)\<rbrakk> \<Longrightarrow>(K +# M = K +# N) \<longleftrightarrow> (M = N)"
 by (auto simp add: multiset_equality)
 
-lemma nat_add_eq_1_cases: "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> (m #+ n = 1) \<longleftrightarrow> (m=1 & n=0) | (m=0 & n=1)"
+lemma nat_add_eq_1_cases: "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> (m #+ n = 1) \<longleftrightarrow> (m=1 \<and> n=0) | (m=0 \<and> n=1)"
 by (induct_tac n) auto
 
 lemma munion_is_single:
      "\<lbrakk>multiset(M); multiset(N)\<rbrakk>
-      \<Longrightarrow> (M +# N = {#a#}) \<longleftrightarrow>  (M={#a#} & N=0) | (M = 0 & N = {#a#})"
+      \<Longrightarrow> (M +# N = {#a#}) \<longleftrightarrow>  (M={#a#} \<and> N=0) | (M = 0 \<and> N = {#a#})"
 apply (simp (no_asm_simp) add: multiset_equality)
 apply safe
 apply simp_all
@@ -465,7 +465,7 @@
 done
 
 lemma msingle_is_union: "\<lbrakk>multiset(M); multiset(N)\<rbrakk>
-  \<Longrightarrow> ({#a#} = M +# N) \<longleftrightarrow> ({#a#} = M  & N=0 | M = 0 & {#a#} = N)"
+  \<Longrightarrow> ({#a#} = M +# N) \<longleftrightarrow> ({#a#} = M  \<and> N=0 | M = 0 \<and> {#a#} = N)"
 apply (subgoal_tac " ({#a#} = M +# N) \<longleftrightarrow> (M +# N = {#a#}) ")
 apply (simp (no_asm_simp) add: munion_is_single)
 apply blast
@@ -663,7 +663,7 @@
 by (auto simp add: mset_of_def MCollect_def multiset_def funrestrict_def)
 
 lemma MCollect_mem_iff [iff]:
-     "x \<in> mset_of({#x \<in> M. P(x)#}) \<longleftrightarrow>  x \<in> mset_of(M) & P(x)"
+     "x \<in> mset_of({#x \<in> M. P(x)#}) \<longleftrightarrow>  x \<in> mset_of(M) \<and> P(x)"
 by (simp add: MCollect_def mset_of_def)
 
 lemma mcount_MCollect [simp]:
@@ -680,8 +680,8 @@
 (* and more algebraic laws on multisets *)
 
 lemma munion_eq_conv_diff: "\<lbrakk>multiset(M); multiset(N)\<rbrakk>
-  \<Longrightarrow>  (M +# {#a#} = N +# {#b#}) \<longleftrightarrow>  (M = N & a = b |
-       M = N -# {#a#} +# {#b#} & N = M -# {#b#} +# {#a#})"
+  \<Longrightarrow>  (M +# {#a#} = N +# {#b#}) \<longleftrightarrow>  (M = N \<and> a = b |
+       M = N -# {#a#} +# {#b#} \<and> N = M -# {#b#} +# {#a#})"
 apply (simp del: mcount_single add: multiset_equality)
 apply (rule iffI, erule_tac [2] disjE, erule_tac [3] conjE)
 apply (case_tac "a=b", auto)
@@ -695,7 +695,7 @@
 
 lemma melem_diff_single:
 "multiset(M) \<Longrightarrow>
-  k \<in> mset_of(M -# {#a#}) \<longleftrightarrow> (k=a & 1 < mcount(M,a)) | (k\<noteq> a & k \<in> mset_of(M))"
+  k \<in> mset_of(M -# {#a#}) \<longleftrightarrow> (k=a \<and> 1 < mcount(M,a)) | (k\<noteq> a \<and> k \<in> mset_of(M))"
 apply (simp add: multiset_def)
 apply (simp add: normalize_def mset_of_def msingle_def mdiff_def mcount_def)
 apply (auto dest: domain_type intro: zero_less_diff [THEN iffD1]
@@ -708,7 +708,7 @@
 lemma munion_eq_conv_exist:
 "\<lbrakk>M \<in> Mult(A); N \<in> Mult(A)\<rbrakk>
   \<Longrightarrow> (M +# {#a#} = N +# {#b#}) \<longleftrightarrow>
-      (M=N & a=b | (\<exists>K \<in> Mult(A). M= K +# {#b#} & N=K +# {#a#}))"
+      (M=N \<and> a=b | (\<exists>K \<in> Mult(A). M= K +# {#b#} \<and> N=K +# {#a#}))"
 by (auto simp add: Mult_iff_multiset melem_diff_single munion_eq_conv_diff)
 
 
@@ -727,9 +727,9 @@
 
 lemma multirel1_iff:
 " <N, M> \<in> multirel1(A, r) \<longleftrightarrow>
-  (\<exists>a. a \<in> A &
-  (\<exists>M0. M0 \<in> Mult(A) & (\<exists>K. K \<in> Mult(A) &
-   M=M0 +# {#a#} & N=M0 +# K & (\<forall>b \<in> mset_of(K). <b,a> \<in> r))))"
+  (\<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))))"
 by (auto simp add: multirel1_def Mult_iff_multiset Bex_def)
 
 
@@ -766,8 +766,8 @@
 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) & N = M +# {#a#}) |
-  (\<exists>K. K \<in> Mult(A) & (\<forall>b \<in> mset_of(K). <b, a> \<in> r) & N = M0 +# K)"
+  (\<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)"
 apply (frule multirel1_type [THEN subsetD])
 apply (simp add: multirel1_iff)
 apply (auto simp add: munion_eq_conv_exist)
@@ -917,8 +917,8 @@
 "<M,N> \<in> multirel(A, r) \<Longrightarrow>
      trans[A](r) \<longrightarrow>
      (\<exists>I J K.
-         I \<in> Mult(A) & J \<in> Mult(A) &  K \<in> Mult(A) &
-         N = I +# J & M = I +# K & J \<noteq> 0 &
+         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))"
 apply (simp add: multirel_def Ball_def Bex_def)
 apply (erule converse_trancl_induct)
@@ -969,7 +969,7 @@
 
 lemma msize_eq_succ_imp_eq_union:
      "\<lbrakk>msize(M)=$# succ(n); M \<in> Mult(A); n \<in> nat\<rbrakk>
-      \<Longrightarrow> \<exists>a N. M = N +# {#a#} & N \<in> Mult(A) & a \<in> A"
+      \<Longrightarrow> \<exists>a N. M = N +# {#a#} \<and> N \<in> Mult(A) \<and> a \<in> A"
 apply (drule msize_eq_succ_imp_elem, auto)
 apply (rule_tac x = a in exI)
 apply (rule_tac x = "M -# {#a#}" in exI)
@@ -983,8 +983,8 @@
 lemma one_step_implies_multirel_lemma [rule_format (no_asm)]:
 "n \<in> nat \<Longrightarrow>
    (\<forall>I J K.
-    I \<in> Mult(A) & J \<in> Mult(A) & K \<in> Mult(A) &
-   (msize(J) = $# n & J \<noteq>0 &  (\<forall>k \<in> mset_of(K).  \<exists>j \<in> mset_of(J). <k, j> \<in> r))
+    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))
     \<longrightarrow> <I +# K, I +# J> \<in> multirel(A, r))"
 apply (simp add: Mult_iff_multiset)
 apply (erule nat_induct, clarify)
@@ -1126,7 +1126,7 @@
 lemma munion_multirel_mono:
      "\<lbrakk><M,K> \<in> multirel(A, r); <N,L> \<in> multirel(A, r)\<rbrakk>
       \<Longrightarrow> <M +# N, K +# L> \<in> multirel(A, r)"
-apply (subgoal_tac "M \<in> Mult(A) & N \<in> Mult(A) & K \<in> Mult(A) & L \<in> Mult(A) ")
+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])
 apply (blast intro: munion_multirel_mono1 multirel_trans munion_multirel_mono2)
 done
@@ -1235,7 +1235,7 @@
 apply (blast intro: mless_trans)
 done
 
-lemma mless_le_iff: "M <# N \<longleftrightarrow> (M <#= N & M \<noteq> N)"
+lemma mless_le_iff: "M <# N \<longleftrightarrow> (M <#= N \<and> M \<noteq> N)"
 by (simp add: mle_def, auto)
 
 (** Monotonicity of mless **)
@@ -1253,7 +1253,7 @@
 lemma munion_less_mono1: "\<lbrakk>M <# N; omultiset(K)\<rbrakk> \<Longrightarrow> M +# K <# N +# K"
 by (force dest: munion_less_mono2 simp add: munion_commute)
 
-lemma mless_imp_omultiset: "M <# N \<Longrightarrow> omultiset(M) & omultiset(N)"
+lemma mless_imp_omultiset: "M <# N \<Longrightarrow> omultiset(M) \<and> omultiset(N)"
 by (auto simp add: mless_def omultiset_def dest: multirel_type [THEN subsetD])
 
 lemma munion_less_mono: "\<lbrakk>M <# K; N <# L\<rbrakk> \<Longrightarrow> M +# N <# K +# L"
@@ -1264,7 +1264,7 @@
 
 (* <#= *)
 
-lemma mle_imp_omultiset: "M <#= N \<Longrightarrow> omultiset(M) & omultiset(N)"
+lemma mle_imp_omultiset: "M <#= N \<Longrightarrow> omultiset(M) \<and> omultiset(N)"
 by (auto simp add: mle_def mless_imp_omultiset)
 
 lemma mle_mono: "\<lbrakk>M <#= K;  N <#= L\<rbrakk> \<Longrightarrow> M +# N <#= K +# L"
@@ -1278,7 +1278,7 @@
 
 lemma empty_leI [simp]: "omultiset(M) \<Longrightarrow> 0 <#= M"
 apply (simp add: mle_def mless_def)
-apply (subgoal_tac "\<exists>i. Ord (i) & M \<in> Mult(field(Memrel(i))) ")
+apply (subgoal_tac "\<exists>i. Ord (i) \<and> M \<in> Mult(field(Memrel(i))) ")
  prefer 2 apply (simp add: omultiset_def)
 apply (case_tac "M=0", simp_all, clarify)
 apply (subgoal_tac "<0 +# 0, 0 +# M> \<in> multirel(field (Memrel(i)), Memrel(i))")
--- a/src/ZF/Induct/Mutil.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Induct/Mutil.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -38,7 +38,7 @@
 
 subsection \<open>Basic properties of evnodd\<close>
 
-lemma evnodd_iff: "<i,j>: evnodd(A,b) \<longleftrightarrow> <i,j>: A & (i#+j) mod 2 = b"
+lemma evnodd_iff: "<i,j>: evnodd(A,b) \<longleftrightarrow> <i,j>: A \<and> (i#+j) mod 2 = b"
   by (unfold evnodd_def) blast
 
 lemma evnodd_subset: "evnodd(A, b) \<subseteq> A"
--- a/src/ZF/Induct/PropLog.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Induct/PropLog.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -1,5 +1,5 @@
 (*  Title:      ZF/Induct/PropLog.thy
-    Author:     Tobias Nipkow & Lawrence C Paulson
+    Author:     Tobias Nipkow \<and> Lawrence C Paulson
     Copyright   1993  University of Cambridge
 *)
 
--- a/src/ZF/Induct/Term.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Induct/Term.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -110,7 +110,7 @@
   \<comment> \<open>\<^term>\<open>map\<close> works correctly on the underlying list of terms.\<close>
   apply (induct set: list)
    apply simp
-  apply (subgoal_tac "rank (a) <i & rank (l) < i")
+  apply (subgoal_tac "rank (a) <i \<and> rank (l) < i")
    apply (simp add: rank_of_Ord)
   apply (simp add: list.con_defs)
   apply (blast dest: rank_rls [THEN lt_trans])
--- a/src/ZF/InfDatatype.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/InfDatatype.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -12,7 +12,7 @@
 
 lemma fun_Vcsucc_lemma:
   assumes f: "f \<in> D -> Vfrom(A,csucc(K))" and DK: "|D| \<le> K" and ICK: "InfCard(K)"
-  shows "\<exists>j. f \<in> D -> Vfrom(A,j) & j < csucc(K)"
+  shows "\<exists>j. f \<in> D -> Vfrom(A,j) \<and> j < csucc(K)"
 proof (rule exI, rule conjI)
   show "f \<in> D \<rightarrow> Vfrom(A, \<Union>z\<in>D. \<mu> i. f`z \<in> Vfrom (A,i))"
     proof (rule Pi_type [OF f])
@@ -46,7 +46,7 @@
 
 lemma subset_Vcsucc:
      "\<lbrakk>D \<subseteq> Vfrom(A,csucc(K));  |D| \<le> K;  InfCard(K)\<rbrakk>
-      \<Longrightarrow> \<exists>j. D \<subseteq> Vfrom(A,j) & j < csucc(K)"
+      \<Longrightarrow> \<exists>j. D \<subseteq> Vfrom(A,j) \<and> j < csucc(K)"
 by (simp add: subset_iff_id fun_Vcsucc_lemma)
 
 (*Version for arbitrary index sets*)
--- a/src/ZF/Int.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Int.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -10,7 +10,7 @@
 definition
   intrel :: i  where
     "intrel \<equiv> {p \<in> (nat*nat)*(nat*nat).
-                \<exists>x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1}"
+                \<exists>x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> \<and> x1#+y2 = x2#+y1}"
 
 definition
   int :: i  where
@@ -34,7 +34,7 @@
 
 definition
   znegative   ::      "i=>o"  where
-    "znegative(z) \<equiv> \<exists>x y. x<y & y\<in>nat & <x,y>\<in>z"
+    "znegative(z) \<equiv> \<exists>x y. x<y \<and> y\<in>nat \<and> <x,y>\<in>z"
 
 definition
   iszero      ::      "i=>o"  where
@@ -52,8 +52,8 @@
   zmagnitude  ::      "i=>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 & ((\<not> znegative(z) & z = $# m) |
-                       (znegative(z) & $- z = $# m))"
+     THE m. m\<in>nat \<and> ((\<not> znegative(z) \<and> z = $# m) |
+                       (znegative(z) \<and> $- z = $# m))"
 
 definition
   raw_zmult   ::      "[i,i]=>i"  where
@@ -99,7 +99,7 @@
 
 lemma intrel_iff [simp]:
     "<<x1,y1>,<x2,y2>>: intrel \<longleftrightarrow>
-     x1\<in>nat & y1\<in>nat & x2\<in>nat & y2\<in>nat & x1#+y2 = x2#+y1"
+     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!]:
--- a/src/ZF/IntDiv.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/IntDiv.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -36,8 +36,8 @@
 definition
   quorem :: "[i,i] => o"  where
     "quorem \<equiv> %<a,b> <q,r>.
-                      a = b$*q $+ r &
-                      (#0$<b & #0$\<le>r & r$<b | \<not>(#0$<b) & b$<r & r $\<le> #0)"
+                      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
@@ -298,7 +298,7 @@
 
 lemma zmult_zless_lemma:
      "\<lbrakk>k \<in> int; m \<in> int; n \<in> int\<rbrakk>
-      \<Longrightarrow> (m$*k $< n$*k) \<longleftrightarrow> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"
+      \<Longrightarrow> (m$*k $< n$*k) \<longleftrightarrow> ((#0 $< k \<and> m$<n) | (k $< #0 \<and> n$<m))"
 apply (case_tac "k = #0")
 apply (auto simp add: neq_iff_zless zmult_zless_mono1 zmult_zless_mono1_neg)
 apply (auto simp add: not_zless_iff_zle
@@ -309,25 +309,25 @@
 done
 
 lemma zmult_zless_cancel2:
-     "(m$*k $< n$*k) \<longleftrightarrow> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"
+     "(m$*k $< n$*k) \<longleftrightarrow> ((#0 $< k \<and> m$<n) | (k $< #0 \<and> n$<m))"
 apply (cut_tac k = "intify (k)" and m = "intify (m)" and n = "intify (n)"
        in zmult_zless_lemma)
 apply auto
 done
 
 lemma zmult_zless_cancel1:
-     "(k$*m $< k$*n) \<longleftrightarrow> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"
+     "(k$*m $< k$*n) \<longleftrightarrow> ((#0 $< k \<and> m$<n) | (k $< #0 \<and> n$<m))"
 by (simp add: zmult_commute [of k] zmult_zless_cancel2)
 
 lemma zmult_zle_cancel2:
-     "(m$*k $\<le> n$*k) \<longleftrightarrow> ((#0 $< k \<longrightarrow> m$\<le>n) & (k $< #0 \<longrightarrow> n$\<le>m))"
+     "(m$*k $\<le> n$*k) \<longleftrightarrow> ((#0 $< k \<longrightarrow> m$\<le>n) \<and> (k $< #0 \<longrightarrow> n$\<le>m))"
 by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel2)
 
 lemma zmult_zle_cancel1:
-     "(k$*m $\<le> k$*n) \<longleftrightarrow> ((#0 $< k \<longrightarrow> m$\<le>n) & (k $< #0 \<longrightarrow> n$\<le>m))"
+     "(k$*m $\<le> k$*n) \<longleftrightarrow> ((#0 $< k \<longrightarrow> m$\<le>n) \<and> (k $< #0 \<longrightarrow> n$\<le>m))"
 by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel1)
 
-lemma int_eq_iff_zle: "\<lbrakk>m \<in> int; n \<in> int\<rbrakk> \<Longrightarrow> m=n \<longleftrightarrow> (m $\<le> n & n $\<le> m)"
+lemma int_eq_iff_zle: "\<lbrakk>m \<in> int; n \<in> int\<rbrakk> \<Longrightarrow> m=n \<longleftrightarrow> (m $\<le> n \<and> n $\<le> m)"
 apply (blast intro: zle_refl zle_anti_sym)
 done
 
@@ -462,7 +462,7 @@
 
 (*FIXME: use intify in integ_of so that we always have @{term"integ_of w \<in> int"}.
     then this rewrite can work for all constants\<And>*)
-lemma intify_eq_0_iff_zle: "intify(m) = #0 \<longleftrightarrow> (m $\<le> #0 & #0 $\<le> m)"
+lemma intify_eq_0_iff_zle: "intify(m) = #0 \<longleftrightarrow> (m $\<le> #0 \<and> #0 $\<le> m)"
   by (simp add: int_eq_iff_zle)
 
 
@@ -482,7 +482,7 @@
 
 lemma int_0_less_lemma:
      "\<lbrakk>x \<in> int; y \<in> int\<rbrakk>
-      \<Longrightarrow> (#0 $< x $* y) \<longleftrightarrow> (#0 $< x & #0 $< y | x $< #0 & y $< #0)"
+      \<Longrightarrow> (#0 $< x $* y) \<longleftrightarrow> (#0 $< x \<and> #0 $< y | x $< #0 \<and> y $< #0)"
 apply (auto simp add: zle_def not_zless_iff_zle zmult_pos zmult_neg)
 apply (rule ccontr)
 apply (rule_tac [2] ccontr)
@@ -496,30 +496,30 @@
 done
 
 lemma int_0_less_mult_iff:
-     "(#0 $< x $* y) \<longleftrightarrow> (#0 $< x & #0 $< y | x $< #0 & y $< #0)"
+     "(#0 $< x $* y) \<longleftrightarrow> (#0 $< x \<and> #0 $< y | x $< #0 \<and> y $< #0)"
 apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_less_lemma)
 apply auto
 done
 
 lemma int_0_le_lemma:
      "\<lbrakk>x \<in> int; y \<in> int\<rbrakk>
-      \<Longrightarrow> (#0 $\<le> x $* y) \<longleftrightarrow> (#0 $\<le> x & #0 $\<le> y | x $\<le> #0 & y $\<le> #0)"
+      \<Longrightarrow> (#0 $\<le> x $* y) \<longleftrightarrow> (#0 $\<le> x \<and> #0 $\<le> y | x $\<le> #0 \<and> y $\<le> #0)"
 by (auto simp add: zle_def not_zless_iff_zle int_0_less_mult_iff)
 
 lemma int_0_le_mult_iff:
-     "(#0 $\<le> x $* y) \<longleftrightarrow> ((#0 $\<le> x & #0 $\<le> y) | (x $\<le> #0 & y $\<le> #0))"
+     "(#0 $\<le> x $* y) \<longleftrightarrow> ((#0 $\<le> x \<and> #0 $\<le> y) | (x $\<le> #0 \<and> y $\<le> #0))"
 apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_le_lemma)
 apply auto
 done
 
 lemma zmult_less_0_iff:
-     "(x $* y $< #0) \<longleftrightarrow> (#0 $< x & y $< #0 | x $< #0 & #0 $< y)"
+     "(x $* y $< #0) \<longleftrightarrow> (#0 $< x \<and> y $< #0 | x $< #0 \<and> #0 $< y)"
 apply (auto simp add: int_0_le_mult_iff not_zle_iff_zless [THEN iff_sym])
 apply (auto dest: zless_not_sym simp add: not_zle_iff_zless)
 done
 
 lemma zmult_le_0_iff:
-     "(x $* y $\<le> #0) \<longleftrightarrow> (#0 $\<le> x & y $\<le> #0 | x $\<le> #0 & #0 $\<le> y)"
+     "(x $* y $\<le> #0) \<longleftrightarrow> (#0 $\<le> x \<and> y $\<le> #0 | x $\<le> #0 \<and> #0 $\<le> y)"
 by (auto dest: zless_not_sym
          simp add: int_0_less_mult_iff not_zless_iff_zle [THEN iff_sym])
 
@@ -778,7 +778,7 @@
 apply auto
 done
 
-lemma pos_mod: "#0 $< b \<Longrightarrow> #0 $\<le> a zmod b & a zmod b $< b"
+lemma pos_mod: "#0 $< b \<Longrightarrow> #0 $\<le> a zmod b \<and> a zmod b $< b"
 apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct)
 apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def)
 apply (blast dest: zle_zless_trans)+
@@ -787,7 +787,7 @@
 lemmas pos_mod_sign = pos_mod [THEN conjunct1]
   and pos_mod_bound = pos_mod [THEN conjunct2]
 
-lemma neg_mod: "b $< #0 \<Longrightarrow> a zmod b $\<le> #0 & b $< a zmod b"
+lemma neg_mod: "b $< #0 \<Longrightarrow> a zmod b $\<le> #0 \<and> b $< a zmod b"
 apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct)
 apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def)
 apply (blast dest: zle_zless_trans)
--- a/src/ZF/List.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/List.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -136,11 +136,11 @@
 (*An elimination rule, for type-checking*)
 inductive_cases ConsE: "Cons(a,l) \<in> list(A)"
 
-lemma Cons_type_iff [simp]: "Cons(a,l) \<in> list(A) \<longleftrightarrow> a \<in> A & l \<in> list(A)"
+lemma Cons_type_iff [simp]: "Cons(a,l) \<in> list(A) \<longleftrightarrow> a \<in> A \<and> l \<in> list(A)"
 by (blast elim: ConsE)
 
 (*Proving freeness results*)
-lemma Cons_iff: "Cons(a,l)=Cons(a',l') \<longleftrightarrow> a=a' & l=l'"
+lemma Cons_iff: "Cons(a,l)=Cons(a',l') \<longleftrightarrow> a=a' \<and> l=l'"
 by auto
 
 lemma Nil_Cons_iff: "\<not> Nil=Cons(a,l)"
@@ -480,7 +480,7 @@
 apply (auto dest: not_lt_imp_le)
 done
 
-lemma lt_min_iff: "\<lbrakk>i \<in> nat; j \<in> nat; k \<in> nat\<rbrakk> \<Longrightarrow> i<min(j,k) \<longleftrightarrow> i<j & i<k"
+lemma lt_min_iff: "\<lbrakk>i \<in> nat; j \<in> nat; k \<in> nat\<rbrakk> \<Longrightarrow> i<min(j,k) \<longleftrightarrow> i<j \<and> i<k"
 apply (unfold min_def)
 apply (auto dest!: not_lt_imp_le intro: lt_trans2 lt_trans)
 done
@@ -530,17 +530,17 @@
 lemma length_greater_0_iff: "xs:list(A) \<Longrightarrow> 0<length(xs) \<longleftrightarrow> xs \<noteq> Nil"
 by (erule list.induct, auto)
 
-lemma length_succ_iff: "xs:list(A) \<Longrightarrow> length(xs)=succ(n) \<longleftrightarrow> (\<exists>y ys. xs=Cons(y, ys) & length(ys)=n)"
+lemma length_succ_iff: "xs:list(A) \<Longrightarrow> length(xs)=succ(n) \<longleftrightarrow> (\<exists>y ys. xs=Cons(y, ys) \<and> length(ys)=n)"
 by (erule list.induct, auto)
 
 (** more theorems about append **)
 
 lemma append_is_Nil_iff [simp]:
-     "xs:list(A) \<Longrightarrow> (xs@ys = Nil) \<longleftrightarrow> (xs=Nil & ys = Nil)"
+     "xs:list(A) \<Longrightarrow> (xs@ys = Nil) \<longleftrightarrow> (xs=Nil \<and> ys = Nil)"
 by (erule list.induct, auto)
 
 lemma append_is_Nil_iff2 [simp]:
-     "xs:list(A) \<Longrightarrow> (Nil = xs@ys) \<longleftrightarrow> (xs=Nil & ys = Nil)"
+     "xs:list(A) \<Longrightarrow> (Nil = xs@ys) \<longleftrightarrow> (xs=Nil \<and> ys = Nil)"
 by (erule list.induct, auto)
 
 lemma append_left_is_self_iff [simp]:
@@ -554,7 +554,7 @@
 (*TOO SLOW as a default simprule!*)
 lemma append_left_is_Nil_iff [rule_format]:
      "\<lbrakk>xs:list(A); ys:list(A); zs:list(A)\<rbrakk> \<Longrightarrow>
-   length(ys)=length(zs) \<longrightarrow> (xs@ys=zs \<longleftrightarrow> (xs=Nil & ys=zs))"
+   length(ys)=length(zs) \<longrightarrow> (xs@ys=zs \<longleftrightarrow> (xs=Nil \<and> ys=zs))"
 apply (erule list.induct)
 apply (auto simp add: length_app)
 done
@@ -562,14 +562,14 @@
 (*TOO SLOW as a default simprule!*)
 lemma append_left_is_Nil_iff2 [rule_format]:
      "\<lbrakk>xs:list(A); ys:list(A); zs:list(A)\<rbrakk> \<Longrightarrow>
-   length(ys)=length(zs) \<longrightarrow> (zs=ys@xs \<longleftrightarrow> (xs=Nil & ys=zs))"
+   length(ys)=length(zs) \<longrightarrow> (zs=ys@xs \<longleftrightarrow> (xs=Nil \<and> ys=zs))"
 apply (erule list.induct)
 apply (auto simp add: length_app)
 done
 
 lemma append_eq_append_iff [rule_format]:
      "xs:list(A) \<Longrightarrow> \<forall>ys \<in> list(A).
-      length(xs)=length(ys) \<longrightarrow> (xs@us = ys@vs) \<longleftrightarrow> (xs=ys & us=vs)"
+      length(xs)=length(ys) \<longrightarrow> (xs@us = ys@vs) \<longleftrightarrow> (xs=ys \<and> us=vs)"
 apply (erule list.induct)
 apply (simp (no_asm_simp))
 apply clarify
@@ -580,7 +580,7 @@
 lemma append_eq_append [rule_format]:
   "xs:list(A) \<Longrightarrow>
    \<forall>ys \<in> list(A). \<forall>us \<in> list(A). \<forall>vs \<in> list(A).
-   length(us) = length(vs) \<longrightarrow> (xs@us = ys@vs) \<longrightarrow> (xs=ys & us=vs)"
+   length(us) = length(vs) \<longrightarrow> (xs@us = ys@vs) \<longrightarrow> (xs=ys \<and> us=vs)"
 apply (induct_tac "xs")
 apply (force simp add: length_app, clarify)
 apply (erule_tac a = ys in list.cases, simp)
@@ -590,7 +590,7 @@
 
 lemma append_eq_append_iff2 [simp]:
  "\<lbrakk>xs:list(A); ys:list(A); us:list(A); vs:list(A); length(us)=length(vs)\<rbrakk>
-  \<Longrightarrow>  xs@us = ys@vs \<longleftrightarrow> (xs=ys & us=vs)"
+  \<Longrightarrow>  xs@us = ys@vs \<longleftrightarrow> (xs=ys \<and> us=vs)"
 apply (rule iffI)
 apply (rule append_eq_append, auto)
 done
@@ -606,7 +606,7 @@
 (* Can also be proved from append_eq_append_iff2,
 but the proof requires two more hypotheses: x \<in> A and y \<in> A *)
 lemma append1_eq_iff [rule_format]:
-     "xs:list(A) \<Longrightarrow> \<forall>ys \<in> list(A). xs@[x] = ys@[y] \<longleftrightarrow> (xs = ys & x=y)"
+     "xs:list(A) \<Longrightarrow> \<forall>ys \<in> list(A). xs@[x] = ys@[y] \<longleftrightarrow> (xs = ys \<and> x=y)"
 apply (erule list.induct)
  apply clarify
  apply (erule list.cases)
@@ -769,7 +769,7 @@
 
 lemma set_of_list_conv_nth:
     "xs:list(A)
-     \<Longrightarrow> set_of_list(xs) = {x \<in> A. \<exists>i\<in>nat. i<length(xs) & x = nth(i,xs)}"
+     \<Longrightarrow> set_of_list(xs) = {x \<in> A. \<exists>i\<in>nat. i<length(xs) \<and> x = nth(i,xs)}"
 apply (induct_tac "xs", simp_all)
 apply (rule equalityI, auto)
 apply (rule_tac x = 0 in bexI, auto)
@@ -968,7 +968,7 @@
      "\<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))
-          & x = nth(i, xs) & y = nth(i, 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)
 
 (** list_update **)
@@ -1192,7 +1192,7 @@
 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 & snd(p) #+ i \<in> A, zip(xs,upt(0,length(xs)))))"
+  map(fst, filter(%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)
--- a/src/ZF/Nat.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Nat.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -18,7 +18,7 @@
 definition
   (*Has an unconditional succ case, which is used in "recursor" below.*)
   nat_case :: "[i, i=>i, i]=>i"  where
-    "nat_case(a,b,k) \<equiv> THE y. k=0 & y=a | (\<exists>x. k=succ(x) & y=b(x))"
+    "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
@@ -244,7 +244,7 @@
 
 lemma split_nat_case:
   "P(nat_case(a,b,k)) \<longleftrightarrow>
-   ((k=0 \<longrightarrow> P(a)) & (\<forall>x. k=succ(x) \<longrightarrow> P(b(x))) & (\<not> quasinat(k) \<longrightarrow> P(0)))"
+   ((k=0 \<longrightarrow> P(a)) \<and> (\<forall>x. k=succ(x) \<longrightarrow> P(b(x))) \<and> (\<not> quasinat(k) \<longrightarrow> P(0)))"
 apply (rule nat_cases [of k])
 apply (auto simp add: non_nat_case)
 done
@@ -291,7 +291,7 @@
 apply (blast intro: lt_trans)
 done
 
-lemma Le_iff [iff]: "<x,y> \<in> Le \<longleftrightarrow> x \<le> y & x \<in> nat & y \<in> nat"
+lemma Le_iff [iff]: "<x,y> \<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 16:51:35 2022 +0100
+++ b/src/ZF/OrdQuant.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -15,7 +15,7 @@
 
 definition
   oex :: "[i, i => o] => o"  where
-    "oex(A, P)  \<equiv> \<exists>x. x<A & P(x)"
+    "oex(A, P)  \<equiv> \<exists>x. x<A \<and> P(x)"
 
 definition
   (* Ordinal Union *)
@@ -44,12 +44,12 @@
 lemma [simp]: "\<not>(\<exists>x<0. P(x))"
 by (simp add: oex_def)
 
-lemma [simp]: "(\<forall>x<succ(i). P(x)) <-> (Ord(i) \<longrightarrow> P(i) & (\<forall>x<i. P(x)))"
+lemma [simp]: "(\<forall>x<succ(i). P(x)) <-> (Ord(i) \<longrightarrow> P(i) \<and> (\<forall>x<i. P(x)))"
 apply (simp add: oall_def le_iff)
 apply (blast intro: lt_Ord2)
 done
 
-lemma [simp]: "(\<exists>x<succ(i). P(x)) <-> (Ord(i) & (P(i) | (\<exists>x<i. P(x))))"
+lemma [simp]: "(\<exists>x<succ(i). P(x)) <-> (Ord(i) \<and> (P(i) | (\<exists>x<i. P(x))))"
 apply (simp add: oex_def le_iff)
 apply (blast intro: lt_Ord2)
 done
@@ -189,7 +189,7 @@
 
 definition
   "rex"      :: "[i=>o, i=>o] => o"  where
-    "rex(M, P) \<equiv> \<exists>x. M(x) & P(x)"
+    "rex(M, P) \<equiv> \<exists>x. M(x) \<and> P(x)"
 
 syntax
   "_rall"     :: "[pttrn, i=>o, o] => o"        (\<open>(3\<forall>_[_]./ _)\<close> 10)
@@ -261,14 +261,14 @@
 declare atomize_rall [symmetric, rulify]
 
 lemma rall_simps1:
-     "(\<forall>x[M]. P(x) & Q)   <-> (\<forall>x[M]. P(x)) & ((\<forall>x[M]. False) | Q)"
+     "(\<forall>x[M]. P(x) \<and> Q)   <-> (\<forall>x[M]. P(x)) \<and> ((\<forall>x[M]. False) | Q)"
      "(\<forall>x[M]. P(x) | Q)   <-> ((\<forall>x[M]. P(x)) | Q)"
      "(\<forall>x[M]. P(x) \<longrightarrow> Q) <-> ((\<exists>x[M]. P(x)) \<longrightarrow> Q)"
      "(\<not>(\<forall>x[M]. P(x))) <-> (\<exists>x[M]. \<not>P(x))"
 by blast+
 
 lemma rall_simps2:
-     "(\<forall>x[M]. P & Q(x))   <-> ((\<forall>x[M]. False) | P) & (\<forall>x[M]. Q(x))"
+     "(\<forall>x[M]. P \<and> Q(x))   <-> ((\<forall>x[M]. False) | P) \<and> (\<forall>x[M]. Q(x))"
      "(\<forall>x[M]. P | Q(x))   <-> (P | (\<forall>x[M]. Q(x)))"
      "(\<forall>x[M]. P \<longrightarrow> Q(x)) <-> (P \<longrightarrow> (\<forall>x[M]. Q(x)))"
 by blast+
@@ -276,19 +276,19 @@
 lemmas rall_simps [simp] = rall_simps1 rall_simps2
 
 lemma rall_conj_distrib:
-    "(\<forall>x[M]. P(x) & Q(x)) <-> ((\<forall>x[M]. P(x)) & (\<forall>x[M]. Q(x)))"
+    "(\<forall>x[M]. P(x) \<and> Q(x)) <-> ((\<forall>x[M]. P(x)) \<and> (\<forall>x[M]. Q(x)))"
 by blast
 
 lemma rex_simps1:
-     "(\<exists>x[M]. P(x) & Q) <-> ((\<exists>x[M]. P(x)) & Q)"
-     "(\<exists>x[M]. P(x) | Q) <-> (\<exists>x[M]. P(x)) | ((\<exists>x[M]. True) & Q)"
-     "(\<exists>x[M]. P(x) \<longrightarrow> Q) <-> ((\<forall>x[M]. P(x)) \<longrightarrow> ((\<exists>x[M]. True) & Q))"
+     "(\<exists>x[M]. P(x) \<and> Q) <-> ((\<exists>x[M]. P(x)) \<and> Q)"
+     "(\<exists>x[M]. P(x) | Q) <-> (\<exists>x[M]. P(x)) | ((\<exists>x[M]. True) \<and> Q)"
+     "(\<exists>x[M]. P(x) \<longrightarrow> Q) <-> ((\<forall>x[M]. P(x)) \<longrightarrow> ((\<exists>x[M]. True) \<and> Q))"
      "(\<not>(\<exists>x[M]. P(x))) <-> (\<forall>x[M]. \<not>P(x))"
 by blast+
 
 lemma rex_simps2:
-     "(\<exists>x[M]. P & Q(x)) <-> (P & (\<exists>x[M]. Q(x)))"
-     "(\<exists>x[M]. P | Q(x)) <-> ((\<exists>x[M]. True) & P) | (\<exists>x[M]. Q(x))"
+     "(\<exists>x[M]. P \<and> Q(x)) <-> (P \<and> (\<exists>x[M]. Q(x)))"
+     "(\<exists>x[M]. P | Q(x)) <-> ((\<exists>x[M]. True) \<and> P) | (\<exists>x[M]. Q(x))"
      "(\<exists>x[M]. P \<longrightarrow> Q(x)) <-> (((\<forall>x[M]. False) | P) \<longrightarrow> (\<exists>x[M]. Q(x)))"
 by blast+
 
@@ -307,10 +307,10 @@
 lemma rex_triv_one_point2 [simp]: "(\<exists>x[M]. a=x) <-> ( M(a))"
 by blast
 
-lemma rex_one_point1 [simp]: "(\<exists>x[M]. x=a & P(x)) <-> ( M(a) & P(a))"
+lemma rex_one_point1 [simp]: "(\<exists>x[M]. x=a \<and> P(x)) <-> ( M(a) \<and> P(a))"
 by blast
 
-lemma rex_one_point2 [simp]: "(\<exists>x[M]. a=x & P(x)) <-> ( M(a) & P(a))"
+lemma rex_one_point2 [simp]: "(\<exists>x[M]. a=x \<and> P(x)) <-> ( M(a) \<and> P(a))"
 by blast
 
 lemma rall_one_point1 [simp]: "(\<forall>x[M]. x=a \<longrightarrow> P(x)) <-> ( M(a) \<longrightarrow> P(a))"
@@ -349,7 +349,7 @@
 
 text \<open>Setting up the one-point-rule simproc\<close>
 
-simproc_setup defined_rex ("\<exists>x[M]. P(x) & Q(x)") = \<open>
+simproc_setup defined_rex ("\<exists>x[M]. P(x) \<and> Q(x)") = \<open>
   fn _ => Quantifier1.rearrange_Bex
     (fn ctxt => unfold_tac ctxt @{thms rex_def})
 \<close>
--- a/src/ZF/Order.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Order.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -17,7 +17,7 @@
 
 definition
   part_ord :: "[i,i]=>o"                (*Strict partial ordering*)  where
-   "part_ord(A,r) \<equiv> irrefl(A,r) & trans[A](r)"
+   "part_ord(A,r) \<equiv> irrefl(A,r) \<and> trans[A](r)"
 
 definition
   linear   :: "[i,i]=>o"                (*Strict total ordering*)  where
@@ -25,7 +25,7 @@
 
 definition
   tot_ord  :: "[i,i]=>o"                (*Strict total ordering*)  where
-   "tot_ord(A,r) \<equiv> part_ord(A,r) & linear(A,r)"
+   "tot_ord(A,r) \<equiv> part_ord(A,r) \<and> linear(A,r)"
 
 definition
   "preorder_on(A, r) \<equiv> refl(A, r) \<and> trans[A](r)"
@@ -41,7 +41,7 @@
 
 definition
   well_ord :: "[i,i]=>o"                (*Well-ordering*)  where
-   "well_ord(A,r) \<equiv> tot_ord(A,r) & wf[A](r)"
+   "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
@@ -64,7 +64,7 @@
 
 definition
   first :: "[i, i, i] => o"  where
-    "first(u, X, R) \<equiv> u \<in> X & (\<forall>v\<in>X. v\<noteq>u \<longrightarrow> <u,v> \<in> R)"
+    "first(u, X, R) \<equiv> u \<in> X \<and> (\<forall>v\<in>X. v\<noteq>u \<longrightarrow> <u,v> \<in> R)"
 
 subsection\<open>Immediate Consequences of the Definitions\<close>
 
@@ -102,7 +102,7 @@
 
 (** Derived rules for pred(A,x,r) **)
 
-lemma pred_iff: "y \<in> pred(A,x,r) \<longleftrightarrow> <y,x>:r & y \<in> A"
+lemma pred_iff: "y \<in> pred(A,x,r) \<longleftrightarrow> <y,x>:r \<and> y \<in> A"
 by (unfold pred_def, blast)
 
 lemmas predI = conjI [THEN pred_iff [THEN iffD2]]
@@ -474,7 +474,7 @@
          f \<in> ord_iso(A,r, B,s);  g \<in> ord_iso(A,r, B,s)\<rbrakk> \<Longrightarrow> f = g"
 apply (rule fun_extension)
 apply (erule ord_iso_is_bij [THEN bij_is_fun])+
-apply (subgoal_tac "f`x \<in> B & g`x \<in> B & linear(B,s)")
+apply (subgoal_tac "f`x \<in> B \<and> g`x \<in> B \<and> linear(B,s)")
  apply (simp add: linear_def)
  apply (blast dest: well_ord_iso_unique_lemma)
 apply (blast intro: ord_iso_is_bij bij_is_fun apply_funtype
@@ -517,7 +517,7 @@
 apply (unfold mono_map_def)
 apply (simp (no_asm_simp) add: ord_iso_map_fun)
 apply safe
-apply (subgoal_tac "x \<in> A & ya:A & y \<in> B & yb:B")
+apply (subgoal_tac "x \<in> A \<and> ya:A \<and> y \<in> B \<and> yb:B")
  apply (simp add: apply_equality [OF _  ord_iso_map_fun])
  apply (unfold ord_iso_map_def)
  apply (blast intro: well_ord_iso_preserving, blast)
--- a/src/ZF/OrderArith.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/OrderArith.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -13,21 +13,21 @@
     "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)> & <x',x>:r)   |
-                    (\<exists>y' y. z = <Inr(y'), Inr(y)> & <y',y>:s)}"
+                    (\<exists>x' x. z = <Inl(x'), Inl(x)> \<and> <x',x>:r)   |
+                    (\<exists>y' y. z = <Inr(y'), Inr(y)> \<and> <y',y>:s)}"
 
 definition
   (*lexicographic product of two relations; underlies ordinal multiplication*)
   rmult   :: "[i,i,i,i]=>i"  where
     "rmult(A,r,B,s) \<equiv>
                 {z: (A*B) * (A*B).
-                    \<exists>x' y' x y. z = <<x',y'>, <x,y>> &
-                       (<x',x>: r | (x'=x & <y',y>: s))}"
+                    \<exists>x' y' x y. z = <<x',y'>, <x,y>> \<and>
+                       (<x',x>: r | (x'=x \<and> <y',y>: 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> & <f`x,f`y>: r}"
+    "rvimage(A,f,r) \<equiv> {z \<in> A*A. \<exists>x y. z = <x,y> \<and> <f`x,f`y>: r}"
 
 definition
   measure :: "[i, i\<Rightarrow>i] \<Rightarrow> i"  where
@@ -39,15 +39,15 @@
 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 & b \<in> B"
+    "<Inl(a), Inr(b)> \<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 & a \<in> A & <a',a>:r"
+    "<Inl(a'), Inl(a)> \<in> radd(A,r,B,s)  \<longleftrightarrow>  a':A \<and> a \<in> A \<and> <a',a>:r"
 by (unfold radd_def, blast)
 
 lemma radd_Inr_iff [iff]:
-    "<Inr(b'), Inr(b)> \<in> radd(A,r,B,s) \<longleftrightarrow>  b':B & b \<in> B & <b',b>:s"
+    "<Inr(b'), Inr(b)> \<in> radd(A,r,B,s) \<longleftrightarrow>  b':B \<and> b \<in> B \<and> <b',b>:s"
 by (unfold radd_def, blast)
 
 lemma radd_Inr_Inl_iff [simp]:
@@ -164,8 +164,8 @@
 
 lemma  rmult_iff [iff]:
     "<<a',b'>, <a,b>> \<in> rmult(A,r,B,s) \<longleftrightarrow>
-            (<a',a>: r  & a':A & a \<in> A & b': B & b \<in> B) |
-            (<b',b>: s  & a'=a & a \<in> A & b': B & b \<in> B)"
+            (<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)"
 
 by (unfold rmult_def, blast)
 
@@ -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 & a \<in> A & b \<in> A"
+lemma rvimage_iff: "<a,b> \<in> rvimage(A,f,r)  \<longleftrightarrow>  <f`a,f`b>: r \<and> a \<in> A \<and> b \<in> A"
 by (unfold rvimage_def, blast)
 
 subsubsection\<open>Type checking\<close>
@@ -361,7 +361,7 @@
 lemma wf_rvimage [intro!]: "wf(r) \<Longrightarrow> wf(rvimage(A,f,r))"
 apply (simp (no_asm_use) add: rvimage_def wf_eq_minimal)
 apply clarify
-apply (subgoal_tac "\<exists>w. w \<in> {w: {f`x. x \<in> Q}. \<exists>x. x \<in> Q & (f`x = w) }")
+apply (subgoal_tac "\<exists>w. w \<in> {w: {f`x. x \<in> Q}. \<exists>x. x \<in> Q \<and> (f`x = w) }")
  apply (erule allE)
  apply (erule impE)
  apply assumption
@@ -440,7 +440,7 @@
 
 
 lemma wf_imp_subset_rvimage:
-     "\<lbrakk>wf(r); r \<subseteq> A*A\<rbrakk> \<Longrightarrow> \<exists>i f. Ord(i) & r \<subseteq> rvimage(A, f, Memrel(i))"
+     "\<lbrakk>wf(r); r \<subseteq> A*A\<rbrakk> \<Longrightarrow> \<exists>i f. Ord(i) \<and> r \<subseteq> rvimage(A, f, Memrel(i))"
 apply (rule_tac x="wftype(r)" in exI)
 apply (rule_tac x="\<lambda>x\<in>A. wfrank(r,x)" in exI)
 apply (simp add: Ord_wftype, clarify)
@@ -450,7 +450,7 @@
 done
 
 theorem wf_iff_subset_rvimage:
-  "relation(r) \<Longrightarrow> wf(r) \<longleftrightarrow> (\<exists>i f A. Ord(i) & r \<subseteq> rvimage(A, f, Memrel(i)))"
+  "relation(r) \<Longrightarrow> wf(r) \<longleftrightarrow> (\<exists>i f A. Ord(i) \<and> r \<subseteq> rvimage(A, f, Memrel(i)))"
 by (blast dest!: relation_field_times_field wf_imp_subset_rvimage
           intro: wf_rvimage_Ord [THEN wf_subset])
 
@@ -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 & y \<in> A & f(x)<f(y)"
+lemma measure_iff [iff]: "<x,y> \<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:
@@ -530,7 +530,7 @@
  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 & u \<in> B(a')) | u \<in> B(a)"
+                       \<Longrightarrow> (\<exists>a'\<in>A. <a',a> \<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)
--- a/src/ZF/OrderType.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/OrderType.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -22,7 +22,7 @@
 definition
   (*alternative definition of ordinal numbers*)
   Ord_alt   :: "i => o"  where
-   "Ord_alt(X) \<equiv> well_ord(X, Memrel(X)) & (\<forall>u\<in>X. u=pred(X, u, Memrel(X)))"
+   "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*)
@@ -600,7 +600,7 @@
                  Union_eq_UN [symmetric] Limit_Union_eq)
 done
 
-lemma oadd_eq_0_iff: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> (i ++ j) = 0 \<longleftrightarrow> i=0 & j=0"
+lemma oadd_eq_0_iff: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> (i ++ j) = 0 \<longleftrightarrow> i=0 \<and> j=0"
 apply (erule trans_induct3 [of j])
 apply (simp_all add: oadd_Limit)
 apply (simp add: Union_empty_iff Limit_def lt_def, blast)
@@ -668,12 +668,12 @@
 done
 
 text\<open>Every ordinal is exceeded by some limit ordinal.\<close>
-lemma Ord_imp_greater_Limit: "Ord(i) \<Longrightarrow> \<exists>k. i<k & Limit(k)"
+lemma Ord_imp_greater_Limit: "Ord(i) \<Longrightarrow> \<exists>k. i<k \<and> Limit(k)"
 apply (rule_tac x="i ++ nat" in exI)
 apply (blast intro: oadd_LimitI  oadd_lt_self  Limit_nat [THEN Limit_has_0])
 done
 
-lemma Ord2_imp_greater_Limit: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> \<exists>k. i<k & j<k & Limit(k)"
+lemma Ord2_imp_greater_Limit: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> \<exists>k. i<k \<and> j<k \<and> Limit(k)"
 apply (insert Ord_Un [of i j, THEN Ord_imp_greater_Limit])
 apply (simp add: Un_least_lt_iff)
 done
@@ -796,7 +796,7 @@
 
 lemma lt_omult:
  "\<lbrakk>Ord(i);  Ord(j);  k<j**i\<rbrakk>
-  \<Longrightarrow> \<exists>j' i'. k = j**i' ++ j' & j'<j & i'<i"
+  \<Longrightarrow> \<exists>j' i'. k = j**i' ++ j' \<and> j'<j \<and> i'<i"
 apply (unfold omult_def)
 apply (simp add: ordertype_pred_unfold well_ord_rmult well_ord_Memrel)
 apply (safe elim!: ltE)
--- a/src/ZF/Ordinal.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Ordinal.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -9,7 +9,7 @@
 
 definition
   Memrel        :: "i=>i"  where
-    "Memrel(A)   \<equiv> {z\<in>A*A . \<exists>x y. z=<x,y> & x\<in>y }"
+    "Memrel(A)   \<equiv> {z\<in>A*A . \<exists>x y. z=<x,y> \<and> x\<in>y }"
 
 definition
   Transset  :: "i=>o"  where
@@ -17,15 +17,15 @@
 
 definition
   Ord  :: "i=>o"  where
-    "Ord(i)      \<equiv> Transset(i) & (\<forall>x\<in>i. Transset(x))"
+    "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
-    "i<j         \<equiv> i\<in>j & Ord(j)"
+    "i<j         \<equiv> i\<in>j \<and> Ord(j)"
 
 definition
   Limit         :: "i=>o"  where
-    "Limit(i)    \<equiv> Ord(i) & 0<i & (\<forall>y. y<i \<longrightarrow> succ(y)<i)"
+    "Limit(i)    \<equiv> Ord(i) \<and> 0<i \<and> (\<forall>y. y<i \<longrightarrow> succ(y)<i)"
 
 abbreviation
   le  (infixl \<open>\<le>\<close> 50) where
@@ -50,11 +50,11 @@
 subsubsection\<open>Consequences of Downwards Closure\<close>
 
 lemma Transset_doubleton_D:
-    "\<lbrakk>Transset(C); {a,b}: C\<rbrakk> \<Longrightarrow> a\<in>C & b\<in>C"
+    "\<lbrakk>Transset(C); {a,b}: C\<rbrakk> \<Longrightarrow> a\<in>C \<and> b\<in>C"
 by (unfold Transset_def, blast)
 
 lemma Transset_Pair_D:
-    "\<lbrakk>Transset(C); <a,b>\<in>C\<rbrakk> \<Longrightarrow> a\<in>C & b\<in>C"
+    "\<lbrakk>Transset(C); <a,b>\<in>C\<rbrakk> \<Longrightarrow> a\<in>C \<and> b\<in>C"
 apply (simp add: Pair_def)
 apply (blast dest: Transset_doubleton_D)
 done
@@ -232,7 +232,7 @@
 
 text\<open>Recall that  \<^term>\<open>i \<le> j\<close>  abbreviates  \<^term>\<open>i<succ(j)\<close> \<And>\<close>
 
-lemma le_iff: "i \<le> j <-> i<j | (i=j & Ord(j))"
+lemma le_iff: "i \<le> j <-> i<j | (i=j \<and> Ord(j))"
 by (unfold lt_def, blast)
 
 (*Equivalently, i<j \<Longrightarrow> i < succ(j)*)
@@ -247,7 +247,7 @@
 lemma le_refl_iff [iff]: "i \<le> i <-> Ord(i)"
 by (simp (no_asm_simp) add: lt_not_refl le_iff)
 
-lemma leCI: "(\<not> (i=j & Ord(j)) \<Longrightarrow> i<j) \<Longrightarrow> i \<le> j"
+lemma leCI: "(\<not> (i=j \<and> Ord(j)) \<Longrightarrow> i<j) \<Longrightarrow> i \<le> j"
 by (simp add: le_iff, blast)
 
 lemma leE:
@@ -267,7 +267,7 @@
 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 & a\<in>A & b\<in>A"
+lemma Memrel_iff [simp]: "<a,b> \<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)"
@@ -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 & b\<in>A"
+    "Transset(A) \<Longrightarrow> <a,b> \<in> Memrel(A) <-> a\<in>b \<and> b\<in>A"
 by (unfold Transset_def, blast)
 
 
@@ -428,10 +428,10 @@
 lemma le_imp_subset: "i \<le> j \<Longrightarrow> i<=j"
 by (blast dest: OrdmemD elim: ltE leE)
 
-lemma le_subset_iff: "j \<le> i <-> j<=i & Ord(i) & Ord(j)"
+lemma le_subset_iff: "j \<le> i <-> j<=i \<and> Ord(i) \<and> Ord(j)"
 by (blast dest: subset_imp_le le_imp_subset elim: ltE)
 
-lemma le_succ_iff: "i \<le> succ(j) <-> i \<le> j | i=succ(j) & Ord(i)"
+lemma le_succ_iff: "i \<le> succ(j) <-> i \<le> j | i=succ(j) \<and> Ord(i)"
 apply (simp (no_asm) add: le_iff)
 apply blast
 done
@@ -476,7 +476,7 @@
 lemma lt_imp_0_lt: "j<i \<Longrightarrow> 0<i"
 by (blast intro: lt_trans1 Ord_0_le [OF lt_Ord])
 
-lemma succ_lt_iff: "succ(i) < j <-> i<j & succ(i) \<noteq> j"
+lemma succ_lt_iff: "succ(i) < j <-> i<j \<and> succ(i) \<noteq> j"
 apply auto
 apply (blast intro: lt_trans le_refl dest: lt_Ord)
 apply (frule lt_Ord)
@@ -506,14 +506,14 @@
 apply (auto simp add: Un_commute le_subset_iff subset_Un_iff lt_Ord)
 done
 
-lemma Un_least_lt_iff: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> i \<union> j < k  <->  i<k & j<k"
+lemma Un_least_lt_iff: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> i \<union> j < k  <->  i<k \<and> j<k"
 apply (safe intro!: Un_least_lt)
 apply (rule_tac [2] Un_upper2_le [THEN lt_trans1])
 apply (rule Un_upper1_le [THEN lt_trans1], auto)
 done
 
 lemma Un_least_mem_iff:
-    "\<lbrakk>Ord(i); Ord(j); Ord(k)\<rbrakk> \<Longrightarrow> i \<union> j \<in> k  <->  i\<in>k & j\<in>k"
+    "\<lbrakk>Ord(i); Ord(j); Ord(k)\<rbrakk> \<Longrightarrow> i \<union> j \<in> k  <->  i\<in>k \<and> j\<in>k"
 apply (insert Un_least_lt_iff [of i j k])
 apply (simp add: lt_def)
 done
@@ -697,7 +697,7 @@
 
 subsubsection\<open>Traditional 3-Way Case Analysis on Ordinals\<close>
 
-lemma Ord_cases_disj: "Ord(i) \<Longrightarrow> i=0 | (\<exists>j. Ord(j) & i=succ(j)) | Limit(i)"
+lemma Ord_cases_disj: "Ord(i) \<Longrightarrow> i=0 | (\<exists>j. Ord(j) \<and> i=succ(j)) | Limit(i)"
 by (blast intro!: non_succ_LimitI Ord_0_lt)
 
 lemma Ord_cases:
--- a/src/ZF/Perm.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Perm.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -16,7 +16,7 @@
   (*composition of relations and functions; NOT Suppes's relative product*)
   comp     :: "[i,i]=>i"      (infixr \<open>O\<close> 60)  where
     "r O s \<equiv> {xz \<in> domain(s)*range(r) .
-               \<exists>x y z. xz=<x,z> & <x,y>:s & <y,z>:r}"
+               \<exists>x y z. xz=<x,z> \<and> <x,y>:s \<and> <y,z>:r}"
 
 definition
   (*the identity function for 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 & y \<in> A"
+lemma id_iff [simp]: "<x,y> \<in> id(A) \<longleftrightarrow> x=y \<and> y \<in> A"
 by auto
 
 
--- a/src/ZF/QPair.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/QPair.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -38,7 +38,7 @@
 
 definition
   qconverse :: "i => i"  where
-    "qconverse(r) \<equiv> {z. w \<in> r, \<exists>x y. w=<x;y> & z=<y;x>}"
+    "qconverse(r) \<equiv> {z. w \<in> r, \<exists>x y. w=<x;y> \<and> z=<y;x>}"
 
 definition
   QSigma    :: "[i, i => i] => i"  where
@@ -77,7 +77,7 @@
 lemma QPair_empty [simp]: "<0;0> = 0"
 by (simp add: QPair_def)
 
-lemma QPair_iff [simp]: "<a;b> = <c;d> \<longleftrightarrow> a=c & b=d"
+lemma QPair_iff [simp]: "<a;b> = <c;d> \<longleftrightarrow> a=c \<and> b=d"
 apply (simp add: QPair_def)
 apply (rule sum_equal_iff)
 done
@@ -263,13 +263,13 @@
 (** <+> is itself injective... who cares?? **)
 
 lemma qsum_iff:
-     "u \<in> A <+> B \<longleftrightarrow> (\<exists>x. x \<in> A & u=QInl(x)) | (\<exists>y. y \<in> B & u=QInr(y))"
+     "u \<in> A <+> B \<longleftrightarrow> (\<exists>x. x \<in> A \<and> u=QInl(x)) | (\<exists>y. y \<in> B \<and> u=QInr(y))"
 by blast
 
-lemma qsum_subset_iff: "A <+> B \<subseteq> C <+> D \<longleftrightarrow> A<=C & B<=D"
+lemma qsum_subset_iff: "A <+> B \<subseteq> C <+> D \<longleftrightarrow> A<=C \<and> B<=D"
 by blast
 
-lemma qsum_equal_iff: "A <+> B = C <+> D \<longleftrightarrow> A=C & B=D"
+lemma qsum_equal_iff: "A <+> B = C <+> D \<longleftrightarrow> A=C \<and> B=D"
 apply (simp (no_asm) add: extension qsum_subset_iff)
 apply blast
 done
--- a/src/ZF/QUniv.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/QUniv.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -27,7 +27,7 @@
 subsection\<open>Properties involving Transset and Sum\<close>
 
 lemma Transset_includes_summands:
-     "\<lbrakk>Transset(C); A+B \<subseteq> C\<rbrakk> \<Longrightarrow> A \<subseteq> C & B \<subseteq> C"
+     "\<lbrakk>Transset(C); A+B \<subseteq> C\<rbrakk> \<Longrightarrow> A \<subseteq> C \<and> B \<subseteq> C"
 apply (simp add: sum_def Un_subset_iff)
 apply (blast dest: Transset_includes_range)
 done
@@ -121,7 +121,7 @@
 
 (*The opposite inclusion*)
 lemma quniv_QPair_D:
-    "<a;b> \<in> quniv(A) \<Longrightarrow> a: quniv(A) & b: quniv(A)"
+    "<a;b> \<in> quniv(A) \<Longrightarrow> a: quniv(A) \<and> b: quniv(A)"
 apply (unfold quniv_def QPair_def)
 apply (rule Transset_includes_summands [THEN conjE])
 apply (rule Transset_eclose [THEN Transset_univ])
@@ -130,7 +130,7 @@
 
 lemmas quniv_QPair_E = quniv_QPair_D [THEN conjE]
 
-lemma quniv_QPair_iff: "<a;b> \<in> quniv(A) \<longleftrightarrow> a: quniv(A) & b: quniv(A)"
+lemma quniv_QPair_iff: "<a;b> \<in> quniv(A) \<longleftrightarrow> a: quniv(A) \<and> b: quniv(A)"
 by (blast intro: QPair_in_quniv dest: quniv_QPair_D)
 
 
--- a/src/ZF/Resid/Confluence.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Resid/Confluence.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -8,12 +8,12 @@
 definition
   confluence    :: "i=>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 & <z,u> \<in> R))"
+       \<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))"
 
 definition
   strip         :: "o"  where
     "strip \<equiv> \<forall>x y. (x =\<Longrightarrow> y) \<longrightarrow> 
-                    (\<forall>z.(x =1=> z) \<longrightarrow> (\<exists>u.(y =1=> u) & (z=\<Longrightarrow>u)))" 
+                    (\<forall>z.(x =1=> z) \<longrightarrow> (\<exists>u.(y =1=> u) \<and> (z=\<Longrightarrow>u)))" 
 
 
 (* ------------------------------------------------------------------------- *)
@@ -103,7 +103,7 @@
 (*      Church_Rosser Theorem                                                *)
 (* ------------------------------------------------------------------------- *)
 
-lemma Church_Rosser: "m<-\<longrightarrow>n \<Longrightarrow> \<exists>p.(m -\<longrightarrow>p) & (n -\<longrightarrow> p)"
+lemma Church_Rosser: "m<-\<longrightarrow>n \<Longrightarrow> \<exists>p.(m -\<longrightarrow>p) \<and> (n -\<longrightarrow> p)"
 apply (erule Sconv.induct)
 apply (erule Sconv1.induct)
 apply (blast intro: red1D1 redD2)
--- a/src/ZF/Resid/Reduction.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Resid/Reduction.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -214,7 +214,7 @@
 (*      Simulation                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-lemma simulation: "m=1=>n \<Longrightarrow> \<exists>v. m|>v = n & m \<sim> v & regular(v)"
+lemma simulation: "m=1=>n \<Longrightarrow> \<exists>v. m|>v = n \<and> m \<sim> v \<and> regular(v)"
 by (erule Spar_red1.induct, force+)
 
 
--- a/src/ZF/Resid/Residuals.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Resid/Residuals.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -199,8 +199,8 @@
 subsection\<open>paving theorem\<close>
 
 lemma paving: "\<lbrakk>w \<sim> u; w \<sim> v; regular(u); regular(v)\<rbrakk>\<Longrightarrow>  
-           \<exists>uv vu. (w|>u) |> vu = (w|>v) |> uv & (w|>u) \<sim> vu \<and>
-             regular(vu) & (w|>v) \<sim> uv \<and> regular(uv)"
+           \<exists>uv vu. (w|>u) |> vu = (w|>v) |> uv \<and> (w|>u) \<sim> vu \<and>
+             regular(vu) \<and> (w|>v) \<sim> uv \<and> regular(uv)"
 apply (subgoal_tac "u \<sim> v")
 apply (safe intro!: exI)
 apply (rule cube)
--- a/src/ZF/Sum.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Sum.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -28,7 +28,7 @@
 subsection\<open>Rules for the \<^term>\<open>Part\<close> Primitive\<close>
 
 lemma Part_iff:
-    "a \<in> Part(A,h) \<longleftrightarrow> a \<in> A & (\<exists>y. a=h(y))"
+    "a \<in> Part(A,h) \<longleftrightarrow> a \<in> A \<and> (\<exists>y. a=h(y))"
 apply (unfold Part_def)
 apply (rule separation)
 done
@@ -106,7 +106,7 @@
 lemma InrD: "Inr(b): A+B \<Longrightarrow> b \<in> B"
 by blast
 
-lemma sum_iff: "u \<in> A+B \<longleftrightarrow> (\<exists>x. x \<in> A & u=Inl(x)) | (\<exists>y. y \<in> B & u=Inr(y))"
+lemma sum_iff: "u \<in> A+B \<longleftrightarrow> (\<exists>x. x \<in> A \<and> u=Inl(x)) | (\<exists>y. y \<in> B \<and> u=Inr(y))"
 by blast
 
 lemma Inl_in_sum_iff [simp]: "(Inl(x) \<in> A+B) \<longleftrightarrow> (x \<in> A)"
@@ -115,10 +115,10 @@
 lemma Inr_in_sum_iff [simp]: "(Inr(y) \<in> A+B) \<longleftrightarrow> (y \<in> B)"
 by auto
 
-lemma sum_subset_iff: "A+B \<subseteq> C+D \<longleftrightarrow> A<=C & B<=D"
+lemma sum_subset_iff: "A+B \<subseteq> C+D \<longleftrightarrow> A<=C \<and> B<=D"
 by blast
 
-lemma sum_equal_iff: "A+B = C+D \<longleftrightarrow> A=C & B=D"
+lemma sum_equal_iff: "A+B = C+D \<longleftrightarrow> A=C \<and> B=D"
 by (simp add: extension sum_subset_iff, blast)
 
 lemma sum_eq_2_times: "A+A = 2*A"
@@ -142,7 +142,7 @@
 
 lemma expand_case: "u \<in> A+B \<Longrightarrow>
         R(case(c,d,u)) \<longleftrightarrow>
-        ((\<forall>x\<in>A. u = Inl(x) \<longrightarrow> R(c(x))) &
+        ((\<forall>x\<in>A. u = Inl(x) \<longrightarrow> R(c(x))) \<and>
         (\<forall>y\<in>B. u = Inr(y) \<longrightarrow> R(d(y))))"
 by auto
 
--- a/src/ZF/Trancl.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Trancl.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -46,7 +46,7 @@
 
 definition
   equiv    :: "[i,i]=>o"  where
-    "equiv(A,r) \<equiv> r \<subseteq> A*A & refl(A,r) & sym(r) & trans(r)"
+    "equiv(A,r) \<equiv> r \<subseteq> A*A \<and> refl(A,r) \<and> sym(r) \<and> trans(r)"
 
 
 subsection\<open>General properties of relations\<close>
@@ -185,7 +185,7 @@
     "\<lbrakk><a,b> \<in> r^*;  (a=b) \<Longrightarrow> P;
         \<And>y.\<lbrakk><a,y> \<in> r^*;   <y,b> \<in> r\<rbrakk> \<Longrightarrow> P\<rbrakk>
      \<Longrightarrow> P"
-apply (subgoal_tac "a = b | (\<exists>y. <a,y> \<in> r^* & <y,b> \<in> r) ")
+apply (subgoal_tac "a = b | (\<exists>y. <a,y> \<in> r^* \<and> <y,b> \<in> r) ")
 (*see HOL/trancl*)
 apply blast
 apply (erule rtrancl_induct, blast+)
@@ -257,7 +257,7 @@
         <a,b> \<in> r \<Longrightarrow> P;
         \<And>y.\<lbrakk><a,y> \<in> r^+; <y,b> \<in> r\<rbrakk> \<Longrightarrow> P
 \<rbrakk> \<Longrightarrow> P"
-apply (subgoal_tac "<a,b> \<in> r | (\<exists>y. <a,y> \<in> r^+ & <y,b> \<in> r) ")
+apply (subgoal_tac "<a,b> \<in> r | (\<exists>y. <a,y> \<in> r^+ \<and> <y,b> \<in> r) ")
 apply blast
 apply (rule compEpair)
 apply (unfold trancl_def, assumption)
--- a/src/ZF/UNITY/AllocBase.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/AllocBase.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -71,12 +71,12 @@
 
 subsection\<open>Various simple lemmas\<close>
 
-lemma Nclients_NbT_gt_0 [simp]: "0 < Nclients & 0 < NbT"
+lemma Nclients_NbT_gt_0 [simp]: "0 < Nclients \<and> 0 < NbT"
 apply (cut_tac Nclients_pos NbT_pos)
 apply (auto intro: Ord_0_lt)
 done
 
-lemma Nclients_NbT_not_0 [simp]: "Nclients \<noteq> 0 & NbT \<noteq> 0"
+lemma Nclients_NbT_not_0 [simp]: "Nclients \<noteq> 0 \<and> NbT \<noteq> 0"
 by (cut_tac Nclients_pos NbT_pos, auto)
 
 lemma Nclients_type [simp,TC]: "Nclients\<in>nat"
@@ -94,7 +94,7 @@
       (\<forall>i\<in>nat. i<n \<longrightarrow> f(i) $\<le> g(i)) \<longrightarrow>  
       setsum(f, n) $\<le> setsum(g,n)"
 apply (induct_tac "n", simp_all)
-apply (subgoal_tac "Finite(x) & x\<notin>x")
+apply (subgoal_tac "Finite(x) \<and> x\<notin>x")
  prefer 2 apply (simp add: nat_into_Finite mem_not_refl, clarify)
 apply (simp (no_asm_simp) add: succ_def)
 apply (subgoal_tac "\<forall>i\<in>nat. i<x\<longrightarrow> f(i) $\<le> g(i) ")
@@ -135,7 +135,7 @@
 done
 
 lemma bag_of_multiset:
-     "l\<in>list(A) \<Longrightarrow> multiset(bag_of(l)) & mset_of(bag_of(l))<=A"
+     "l\<in>list(A) \<Longrightarrow> multiset(bag_of(l)) \<and> mset_of(bag_of(l))<=A"
 apply (drule bag_of_type)
 apply (auto simp add: Mult_iff_multiset)
 done
@@ -279,7 +279,7 @@
 
 lemma all_distinct_Cons [simp]: 
     "all_distinct(Cons(a,l)) \<longleftrightarrow>  
-      (a\<in>set_of_list(l) \<longrightarrow> False) & (a \<notin> set_of_list(l) \<longrightarrow> all_distinct(l))"
+      (a\<in>set_of_list(l) \<longrightarrow> False) \<and> (a \<notin> set_of_list(l) \<longrightarrow> all_distinct(l))"
 apply (unfold all_distinct_def)
 apply (auto elim: list.cases)
 done
--- a/src/ZF/UNITY/AllocImpl.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/AllocImpl.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -18,31 +18,31 @@
 
 axiomatization where
   alloc_type_assumes [simp]:
-  "type_of(NbR) = nat & type_of(available_tok)=nat" and
+  "type_of(NbR) = nat \<and> type_of(available_tok)=nat" and
 
   alloc_default_val_assumes [simp]:
-  "default_val(NbR)  = 0 & default_val(available_tok)=0"
+  "default_val(NbR)  = 0 \<and> default_val(available_tok)=0"
 
 definition
   "alloc_giv_act \<equiv>
        {<s, t> \<in> state*state.
-        \<exists>k. k = length(s`giv) &
+        \<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)) &
-            k < length(s`ask) & nth(k, s`ask) \<le> s`available_tok}"
+                  available_tok := s`available_tok #- nth(k, s`ask)) \<and>
+            k < length(s`ask) \<and> nth(k, s`ask) \<le> s`available_tok}"
 
 definition
   "alloc_rel_act \<equiv>
        {<s, t> \<in> state*state.
         t = s(available_tok := s`available_tok #+ nth(s`NbR, s`rel),
-              NbR := succ(s`NbR)) &
+              NbR := succ(s`NbR)) \<and>
         s`NbR < length(s`rel)}"
 
 definition
   (*The initial condition s`giv=[] is missing from the
     original definition: S. O. Ehmety *)
   "alloc_prog \<equiv>
-       mk_program({s:state. s`available_tok=NbT & s`NbR=0 & s`giv=Nil},
+       mk_program({s:state. s`available_tok=NbT \<and> s`NbR=0 \<and> s`giv=Nil},
                   {alloc_giv_act, alloc_rel_act},
                   \<Union>G \<in> preserves(lift(available_tok)) \<inter>
                         preserves(lift(NbR)) \<inter>
@@ -74,8 +74,8 @@
 
 lemma alloc_prog_ok_iff:
 "\<forall>G \<in> program. (alloc_prog ok G) \<longleftrightarrow>
-     (G \<in> preserves(lift(giv)) & G \<in> preserves(lift(available_tok)) &
-       G \<in> preserves(lift(NbR)) &  alloc_prog \<in> Allowed(G))"
+     (G \<in> preserves(lift(giv)) \<and> G \<in> preserves(lift(available_tok)) \<and>
+       G \<in> preserves(lift(NbR)) \<and>  alloc_prog \<in> Allowed(G))"
 by (auto simp add: ok_iff_Allowed alloc_prog_def [THEN def_prg_Allowed])
 
 
@@ -151,7 +151,7 @@
 apply (subgoal_tac "G \<in> preserves (fun_pair (lift (available_tok), fun_pair (lift (NbR), lift (giv))))")
 apply (rotate_tac -1)
 apply (cut_tac A = "nat * nat * list(nat)"
-             and P = "%<m,n,l> y. n \<le> length(y) &
+             and P = "%<m,n,l> y. n \<le> length(y) \<and>
                                   m #+ tokens(l) = NbT #+ tokens(take(n,y))"
              and g = "lift(rel)" and F = alloc_prog
        in stable_Join_Stable)
@@ -440,7 +440,7 @@
 by auto
 
 (*needed?*)
-lemma single_state_Diff_eq [simp]: "{s}-{x \<in> state. P(x)} = (if s\<in>state & P(s) then 0 else {s})"
+lemma single_state_Diff_eq [simp]: "{s}-{x \<in> state. P(x)} = (if s\<in>state \<and> P(s) then 0 else {s})"
 by auto
 
 
@@ -494,9 +494,9 @@
 lemma (in alloc_progress) length_giv_disj:
      "\<lbrakk>k \<in> tokbag; n \<in> nat\<rbrakk>
       \<Longrightarrow> alloc_prog \<squnion> G \<in>
-            {s\<in>state. length(s`giv) = n & tokens(s`giv) = k}
+            {s\<in>state. length(s`giv) = n \<and> tokens(s`giv) = k}
             \<longmapsto>w
-              {s\<in>state. (length(s`giv) = n & tokens(s`giv) = k &
+              {s\<in>state. (length(s`giv) = n \<and> tokens(s`giv) = k \<and>
                          k \<le> tokens(take(s`NbR, s`rel))) | n < length(s`giv)}"
 apply (rule single_LeadsTo_I, safe)
 apply (rule_tac a1 = "s`giv" in Increasing_imp_Stable [THEN PSP_StableI])
@@ -514,9 +514,9 @@
 lemma (in alloc_progress) length_giv_disj2:
      "\<lbrakk>k \<in> tokbag; n \<in> nat\<rbrakk>
       \<Longrightarrow> alloc_prog \<squnion> G \<in>
-            {s\<in>state. length(s`giv) = n & tokens(s`giv) = k}
+            {s\<in>state. length(s`giv) = n \<and> tokens(s`giv) = k}
             \<longmapsto>w
-              {s\<in>state. (length(s`giv) = n & NbT \<le> s`available_tok) |
+              {s\<in>state. (length(s`giv) = n \<and> NbT \<le> s`available_tok) |
                         n < length(s`giv)}"
 apply (rule LeadsTo_weaken_R)
 apply (rule Always_LeadsToD [OF alloc_prog_Always_lemma length_giv_disj], auto)
@@ -529,7 +529,7 @@
       \<Longrightarrow> alloc_prog \<squnion> G \<in>
             {s\<in>state. length(s`giv) = n}
             \<longmapsto>w
-              {s\<in>state. (length(s`giv) = n & NbT \<le> s`available_tok) |
+              {s\<in>state. (length(s`giv) = n \<and> NbT \<le> s`available_tok) |
                         n < length(s`giv)}"
 apply (rule LeadsTo_weaken_L)
 apply (rule_tac I = nat in LeadsTo_UN)
@@ -542,9 +542,9 @@
 lemma (in alloc_progress) length_ask_giv:
  "\<lbrakk>k \<in> nat;  n < k\<rbrakk>
   \<Longrightarrow> alloc_prog \<squnion> G \<in>
-        {s\<in>state. length(s`ask) = k & length(s`giv) = n}
+        {s\<in>state. length(s`ask) = k \<and> length(s`giv) = n}
         \<longmapsto>w
-          {s\<in>state. (NbT \<le> s`available_tok & length(s`giv) < length(s`ask) &
+          {s\<in>state. (NbT \<le> s`available_tok \<and> length(s`giv) < length(s`ask) \<and>
                      length(s`giv) = n) |
                     n < length(s`giv)}"
 apply (rule single_LeadsTo_I, safe)
@@ -565,10 +565,10 @@
 lemma (in alloc_progress) length_ask_giv2:
      "\<lbrakk>k \<in> nat;  n < k\<rbrakk>
       \<Longrightarrow> alloc_prog \<squnion> G \<in>
-            {s\<in>state. length(s`ask) = k & length(s`giv) = n}
+            {s\<in>state. length(s`ask) = k \<and> length(s`giv) = n}
             \<longmapsto>w
-              {s\<in>state. (nth(length(s`giv), s`ask) \<le> s`available_tok &
-                         length(s`giv) < length(s`ask) & length(s`giv) = n) |
+              {s\<in>state. (nth(length(s`giv), s`ask) \<le> s`available_tok \<and>
+                         length(s`giv) < length(s`ask) \<and> length(s`giv) = n) |
                         n < length(s`giv)}"
 apply (rule LeadsTo_weaken_R)
 apply (rule Always_LeadsToD [OF safety length_ask_giv], assumption+, clarify)
@@ -582,7 +582,7 @@
 lemma (in alloc_progress) extend_giv:
      "\<lbrakk>k \<in> nat;  n < k\<rbrakk>
       \<Longrightarrow> alloc_prog \<squnion> G \<in>
-            {s\<in>state. length(s`ask) = k & length(s`giv) = n}
+            {s\<in>state. length(s`ask) = k \<and> length(s`giv) = n}
             \<longmapsto>w {s\<in>state. n < length(s`giv)}"
 apply (rule LeadsTo_Un_duplicate)
 apply (rule LeadsTo_cancel1)
--- a/src/ZF/UNITY/ClientImpl.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/ClientImpl.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -14,11 +14,11 @@
 
 axiomatization where
   type_assumes:
-  "type_of(ask) = list(tokbag) & type_of(giv) = list(tokbag) &
-   type_of(rel) = list(tokbag) & type_of(tok) = nat" and
+  "type_of(ask) = list(tokbag) \<and> type_of(giv) = list(tokbag) \<and>
+   type_of(rel) = list(tokbag) \<and> type_of(tok) = nat" and
   default_val_assumes:
-  "default_val(ask) = Nil & default_val(giv) = Nil &
-   default_val(rel) = Nil & default_val(tok) = 0"
+  "default_val(ask) = Nil \<and> default_val(giv) = Nil \<and>
+   default_val(rel) = Nil \<and> default_val(tok) = 0"
 
 
 (*Array indexing is translated to list indexing as A[n] \<equiv> nth(n-1,A). *)
@@ -27,9 +27,9 @@
  (** Release some client_tokens **)
     "client_rel_act \<equiv>
      {<s,t> \<in> state*state.
-      \<exists>nrel \<in> nat. nrel = length(s`rel) &
-                   t = s(rel:=(s`rel)@[nth(nrel, s`giv)]) &
-                   nrel < length(s`giv) &
+      \<exists>nrel \<in> nat. nrel = length(s`rel) \<and>
+                   t = s(rel:=(s`rel)@[nth(nrel, s`giv)]) \<and>
+                   nrel < length(s`giv) \<and>
                    nth(nrel, s`ask) \<le> nth(nrel, s`giv)}"
 
   (** Choose a new token requirement **)
@@ -44,8 +44,8 @@
 
 definition
   "client_prog \<equiv>
-   mk_program({s \<in> state. s`tok \<le> NbT & s`giv = Nil &
-                       s`ask = Nil & s`rel = Nil},
+   mk_program({s \<in> state. s`tok \<le> NbT \<and> s`giv = Nil \<and>
+                       s`ask = Nil \<and> s`rel = Nil},
                     {client_rel_act, client_tok_act, client_ask_act},
                    \<Union>G \<in> preserves(lift(rel)) Int
                          preserves(lift(ask)) Int
@@ -92,8 +92,8 @@
 
 lemma client_prog_ok_iff:
   "\<forall>G \<in> program. (client_prog ok G) \<longleftrightarrow>
-   (G \<in> preserves(lift(rel)) & G \<in> preserves(lift(ask)) &
-    G \<in> preserves(lift(tok)) &  client_prog \<in> Allowed(G))"
+   (G \<in> preserves(lift(rel)) \<and> G \<in> preserves(lift(ask)) \<and>
+    G \<in> preserves(lift(tok)) \<and>  client_prog \<in> Allowed(G))"
 by (auto simp add: ok_iff_Allowed client_prog_def [THEN def_prg_Allowed])
 
 lemma client_prog_preserves:
@@ -186,8 +186,8 @@
 
 lemma transient_lemma:
 "client_prog \<in>
-  transient({s \<in> state. s`rel = k & <k, h> \<in> strict_prefix(nat)
-   & <h, s`giv> \<in> prefix(nat) & h pfixGe s`ask})"
+  transient({s \<in> state. s`rel = k \<and> <k, h> \<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])
 apply (simp (no_asm) add: client_rel_act_def [THEN def_act_eq, THEN act_subset])
@@ -209,7 +209,7 @@
 done
 
 lemma strict_prefix_is_prefix:
-    "<xs, ys> \<in> strict_prefix(A) \<longleftrightarrow>  <xs, ys> \<in> prefix(A) & xs\<noteq>ys"
+    "<xs, ys> \<in> strict_prefix(A) \<longleftrightarrow>  <xs, ys> \<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,11 +217,11 @@
 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 & <k,h> \<in> strict_prefix(nat)
-   & <h, s`giv> \<in> prefix(nat) & h pfixGe s`ask}
+  {s \<in> state. s`rel = k \<and> <k,h> \<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)
-                          & <s`rel, s`giv> \<in> prefix(nat) &
-                                  <h, s`giv> \<in> prefix(nat) &
+                          \<and> <s`rel, s`giv> \<in> prefix(nat) \<and>
+                                  <h, s`giv> \<in> prefix(nat) \<and>
                 h pfixGe s`ask}"
 apply (rule single_LeadsTo_I)
  prefer 2 apply simp
@@ -247,7 +247,7 @@
 "\<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, h> \<in> strict_prefix(nat)
-           & <h, s`giv> \<in> prefix(nat) & h pfixGe s`ask}
+           \<and> <h, s`giv> \<in> prefix(nat) \<and> h pfixGe s`ask}
                       \<longmapsto>w {s \<in> state. <h, s`rel> \<in> prefix(nat)}"
 apply (rule_tac f = "\<lambda>x \<in> state. length(h) #- length(x`rel)"
        in LessThan_induct)
@@ -275,7 +275,7 @@
 lemma progress_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. <h, s`giv> \<in> prefix(nat) & h pfixGe s`ask}
+       \<in> {s \<in> state. <h, s`giv> \<in> prefix(nat) \<and> h pfixGe s`ask}
          \<longmapsto>w  {s \<in> state. <h, s`rel> \<in> prefix(nat)}"
 apply (rule client_prog_Join_Always_rel_le_giv [THEN Always_LeadsToI],
        assumption)
@@ -290,7 +290,7 @@
 (*Progress property: all tokens that are given will be released*)
 lemma client_prog_progress:
 "client_prog \<in> Incr(lift(giv))  guarantees
-      (\<Inter>h \<in> list(nat). {s \<in> state. <h, s`giv> \<in> prefix(nat) &
+      (\<Inter>h \<in> list(nat). {s \<in> state. <h, s`giv> \<in> prefix(nat) \<and>
               h pfixGe s`ask} \<longmapsto>w {s \<in> state. <h, s`rel> \<in> prefix(nat)})"
 apply (rule guaranteesI)
 apply (blast intro: progress_lemma, auto)
--- a/src/ZF/UNITY/Comp.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/Comp.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -23,16 +23,16 @@
 
 definition
   strict_component :: "[i,i]=>o" (infixl \<open>strict'_component\<close> 65)  where
-  "F strict_component H \<equiv> F component H & F\<noteq>H"
+  "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
-  "F component_of H  \<equiv> (\<exists>G. F ok G & F \<squnion> G = H)"
+  "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
-  "F strict_component_of H  \<equiv> F component_of H  & F\<noteq>H"
+  "F strict_component_of H  \<equiv> F component_of H  \<and> F\<noteq>H"
 
 definition
   (* Preserves a state function f, in particular a variable *)
@@ -61,7 +61,7 @@
 
 lemma component_eq_subset: 
      "G \<in> program \<Longrightarrow> (F component G) \<longleftrightarrow>  
-   (Init(G) \<subseteq> Init(F) & Acts(F) \<subseteq> Acts(G) & AllowedActs(G) \<subseteq> AllowedActs(F))"
+   (Init(G) \<subseteq> Init(F) \<and> Acts(F) \<subseteq> Acts(G) \<and> AllowedActs(G) \<subseteq> AllowedActs(F))"
 apply (unfold component_def, auto)
 apply (rule exI)
 apply (rule program_equalityI, auto)
@@ -121,7 +121,7 @@
 done
 
 lemma component_antisym:
-     "\<lbrakk>F \<in> program; G \<in> program\<rbrakk> \<Longrightarrow>(F component G & G  component F) \<longrightarrow> F = G"
+     "\<lbrakk>F \<in> program; G \<in> program\<rbrakk> \<Longrightarrow>(F component G \<and> G  component F) \<longrightarrow> F = G"
 apply (simp (no_asm_simp) add: component_eq_subset)
 apply clarify
 apply (rule program_equalityI, auto)
@@ -129,7 +129,7 @@
 
 lemma Join_component_iff:
      "H \<in> program \<Longrightarrow> 
-      ((F \<squnion> G) component H) \<longleftrightarrow> (F component H & G component H)"
+      ((F \<squnion> G) component H) \<longleftrightarrow> (F component H \<and> G component H)"
 apply (simp (no_asm_simp) add: component_eq_subset)
 apply blast
 done
@@ -158,13 +158,13 @@
 lemma preserves_imp_eq: 
      "\<lbrakk>F \<in> preserves(f);  act \<in> Acts(F);  <s,t> \<in> act\<rbrakk> \<Longrightarrow> f(s) = f(t)"
 apply (unfold preserves_def stable_def constrains_def)
-apply (subgoal_tac "s \<in> state & t \<in> state")
+apply (subgoal_tac "s \<in> state \<and> t \<in> state")
  prefer 2 apply (blast dest!: Acts_type [THEN subsetD], force) 
 done
 
 lemma Join_preserves [iff]: 
 "(F \<squnion> G \<in> preserves(v)) \<longleftrightarrow>   
-      (programify(F) \<in> preserves(v) & programify(G) \<in> preserves(v))"
+      (programify(F) \<in> preserves(v) \<and> programify(G) \<in> preserves(v))"
 by (auto simp add: preserves_def INT_iff)
  
 lemma JOIN_preserves [iff]:
@@ -297,7 +297,7 @@
 lemma Increasing_preserves_Stable:
      "\<lbrakk>F \<in> stable({s \<in> state. <f(s), g(s)>:r});  G \<in> preserves(f);    
          F \<squnion> G \<in> Increasing(A, r, g);  
-         \<forall>x \<in> state. f(x):A & g(x):A\<rbrakk>      
+         \<forall>x \<in> state. f(x):A \<and> g(x):A\<rbrakk>      
       \<Longrightarrow> F \<squnion> G \<in> Stable({s \<in> state. <f(s), g(s)>:r})"
 apply (auto simp add: stable_def Stable_def Increasing_def Constrains_def all_conj_distrib)
 apply (simp_all add: constrains_type [THEN subsetD] preserves_type [THEN subsetD])
--- a/src/ZF/UNITY/Constrains.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/Constrains.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -344,7 +344,7 @@
 apply (frule Stable_type [THEN subsetD], auto)
 done
 
-lemma AlwaysD: "F \<in> Always(A) \<Longrightarrow> Init(F)<=A & F \<in> Stable(A)"
+lemma AlwaysD: "F \<in> Always(A) \<Longrightarrow> Init(F)<=A \<and> F \<in> Stable(A)"
 by (simp add: Always_def initially_def)
 
 lemmas AlwaysE = AlwaysD [THEN conjE]
--- a/src/ZF/UNITY/Distributor.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/Distributor.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -22,7 +22,7 @@
          (\<Inter>n \<in> Nclients.
           lift(Out(n))
               Fols
-          (%s. sublist(s`In, {k \<in> nat. k<length(s`iIn) & nth(k, s`iIn) = n}))
+          (%s. sublist(s`In, {k \<in> nat. k<length(s`iIn) \<and> nth(k, s`iIn) = n}))
           Wrt prefix(A)/list(A))"
 
 definition
@@ -43,13 +43,13 @@
     and A   \<comment> \<open>the type of items being distributed\<close>
     and D
  assumes
-     var_assumes [simp]:  "In \<in> var & iIn \<in> var & (\<forall>n. Out(n):var)"
+     var_assumes [simp]:  "In \<in> var \<and> iIn \<in> var \<and> (\<forall>n. Out(n):var)"
  and all_distinct_vars:   "\<forall>n. all_distinct([In, iIn, Out(n)])"
- and type_assumes [simp]: "type_of(In)=list(A) &  type_of(iIn)=list(nat) &
+ and type_assumes [simp]: "type_of(In)=list(A) \<and>  type_of(iIn)=list(nat) \<and>
                           (\<forall>n. type_of(Out(n))=list(A))"
  and default_val_assumes [simp]:
-                         "default_val(In)=Nil &
-                          default_val(iIn)=Nil &
+                         "default_val(In)=Nil \<and>
+                          default_val(iIn)=Nil \<and>
                           (\<forall>n. default_val(Out(n))=Nil)"
  and distr_spec:  "D \<in> distr_spec(A, In, iIn, Out)"
 
@@ -79,7 +79,7 @@
 
 lemma (in distr) D_ok_iff:
      "G \<in> program \<Longrightarrow>
-        D ok G \<longleftrightarrow> ((\<forall>n \<in> nat. G \<in> preserves(lift(Out(n)))) & D \<in> Allowed(G))"
+        D ok G \<longleftrightarrow> ((\<forall>n \<in> nat. G \<in> preserves(lift(Out(n)))) \<and> D \<in> Allowed(G))"
 apply (cut_tac distr_spec)
 apply (auto simp add: INT_iff distr_spec_def distr_allowed_acts_def
                       Allowed_def ok_iff_Allowed)
@@ -106,7 +106,7 @@
           \<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,
-                       {k \<in> nat. k < length(s`iIn) &
+                       {k \<in> nat. k < length(s`iIn) \<and>
                           nth(k, s`iIn)= n})),
                          Nclients, A) =
               bag_of(sublist(s`In, length(s`iIn)))})"
@@ -122,7 +122,7 @@
  apply (simp (no_asm_simp))
 apply (simp add: set_of_list_conv_nth [of _ nat])
 apply (subgoal_tac
-       "(\<Union>i \<in> Nclients. {k\<in>nat. k < length(s`iIn) & nth(k, s`iIn) = i}) =
+       "(\<Union>i \<in> Nclients. {k\<in>nat. k < length(s`iIn) \<and> nth(k, s`iIn) = i}) =
         length(s`iIn) ")
 apply (simp (no_asm_simp))
 apply (rule equalityI)
--- a/src/ZF/UNITY/Follows.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/Follows.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -51,7 +51,7 @@
 
 
 lemma subset_Always_comp:
-"\<lbrakk>mono1(A, r, B, s, h); \<forall>x \<in> state. f(x):A & g(x):A\<rbrakk> \<Longrightarrow>
+"\<lbrakk>mono1(A, r, B, s, h); \<forall>x \<in> state. f(x):A \<and> g(x):A\<rbrakk> \<Longrightarrow>
    Always({x \<in> state. <f(x), g(x)> \<in> r})<=Always({x \<in> state. <(h comp f)(x), (h comp g)(x)> \<in> s})"
 apply (unfold mono1_def metacomp_def)
 apply (auto simp add: Always_eq_includes_reachable)
@@ -59,7 +59,7 @@
 
 lemma imp_Always_comp:
 "\<lbrakk>F \<in> Always({x \<in> state. <f(x), g(x)> \<in> r});
-    mono1(A, r, B, s, h); \<forall>x \<in> state. f(x):A & g(x):A\<rbrakk> \<Longrightarrow>
+    mono1(A, r, B, s, h); \<forall>x \<in> state. f(x):A \<and> g(x):A\<rbrakk> \<Longrightarrow>
     F \<in> Always({x \<in> state. <(h comp f)(x), (h comp g)(x)> \<in> s})"
 by (blast intro: subset_Always_comp [THEN subsetD])
 
@@ -68,7 +68,7 @@
 "\<lbrakk>F \<in> Always({x \<in> state. <f1(x), f(x)> \<in> r});
     F \<in> Always({x \<in> state. <g1(x), g(x)> \<in> s});
     mono2(A, r, B, s, C, t, h);
-    \<forall>x \<in> state. f1(x):A & f(x):A & g1(x):B & g(x):B\<rbrakk>
+    \<forall>x \<in> state. f1(x):A \<and> f(x):A \<and> g1(x):B \<and> g(x):B\<rbrakk>
   \<Longrightarrow> F \<in> Always({x \<in> state. <h(f1(x), g1(x)), h(f(x), g(x))> \<in> t})"
 apply (auto simp add: Always_eq_includes_reachable mono2_def)
 apply (auto dest!: subsetD)
@@ -78,7 +78,7 @@
 
 lemma subset_LeadsTo_comp:
 "\<lbrakk>mono1(A, r, B, s, h); refl(A,r); trans[B](s);
-        \<forall>x \<in> state. f(x):A & g(x):A\<rbrakk> \<Longrightarrow>
+        \<forall>x \<in> state. f(x):A \<and> g(x):A\<rbrakk> \<Longrightarrow>
   (\<Inter>j \<in> A. {s \<in> state. <j, g(s)> \<in> r} \<longmapsto>w {s \<in> state. <j,f(s)> \<in> r}) <=
  (\<Inter>k \<in> B. {x \<in> state. <k, (h comp g)(x)> \<in> s} \<longmapsto>w {x \<in> state. <k, (h comp f)(x)> \<in> s})"
 
@@ -99,7 +99,7 @@
 lemma imp_LeadsTo_comp:
 "\<lbrakk>F:(\<Inter>j \<in> A. {s \<in> state. <j, g(s)> \<in> r} \<longmapsto>w {s \<in> state. <j,f(s)> \<in> r});
     mono1(A, r, B, s, h); refl(A,r); trans[B](s);
-    \<forall>x \<in> state. f(x):A & g(x):A\<rbrakk> \<Longrightarrow>
+    \<forall>x \<in> state. f(x):A \<and> g(x):A\<rbrakk> \<Longrightarrow>
    F:(\<Inter>k \<in> B. {x \<in> state. <k, (h comp g)(x)> \<in> s} \<longmapsto>w {x \<in> state. <k, (h comp f)(x)> \<in> s})"
 apply (rule subset_LeadsTo_comp [THEN subsetD], auto)
 done
@@ -108,7 +108,7 @@
 "\<lbrakk>F \<in> Increasing(B, s, g);
   \<forall>j \<in> A. F: {s \<in> state. <j, f(s)> \<in> r} \<longmapsto>w {s \<in> state. <j,f1(s)> \<in> r};
   mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t);
-  \<forall>x \<in> state. f1(x):A & f(x):A & g(x):B; k \<in> C\<rbrakk> \<Longrightarrow>
+  \<forall>x \<in> state. f1(x):A \<and> f(x):A \<and> g(x):B; k \<in> C\<rbrakk> \<Longrightarrow>
   F:{x \<in> state. <k, h(f(x), g(x))> \<in> t} \<longmapsto>w {x \<in> state. <k, h(f1(x), g(x))> \<in> t}"
 apply (unfold mono2_def Increasing_def)
 apply (rule single_LeadsTo_I, auto)
@@ -131,7 +131,7 @@
 "\<lbrakk>F \<in> Increasing(A, r, f);
   \<forall>j \<in> B. F: {x \<in> state. <j, g(x)> \<in> s} \<longmapsto>w {x \<in> state. <j,g1(x)> \<in> s};
   mono2(A, r, B, s, C, t, h); refl(A,r); refl(B, s); trans[C](t);
-  \<forall>x \<in> state. f(x):A & g1(x):B & g(x):B; k \<in> C\<rbrakk> \<Longrightarrow>
+  \<forall>x \<in> state. f(x):A \<and> g1(x):B \<and> g(x):B; k \<in> C\<rbrakk> \<Longrightarrow>
   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)
@@ -156,7 +156,7 @@
   \<forall>j \<in> A. F: {s \<in> state. <j, f(s)> \<in> r} \<longmapsto>w {s \<in> state. <j,f1(s)> \<in> r};
   \<forall>j \<in> B. F: {x \<in> state. <j, g(x)> \<in> s} \<longmapsto>w {x \<in> state. <j,g1(x)> \<in> s};
   mono2(A, r, B, s, C, t, h); refl(A,r); refl(B, s); trans[C](t);
-  \<forall>x \<in> state. f(x):A & g1(x):B & f1(x):A &g(x):B; k \<in> C\<rbrakk>
+  \<forall>x \<in> state. f(x):A \<and> g1(x):B \<and> f1(x):A \<and>g(x):B; k \<in> C\<rbrakk>
   \<Longrightarrow> F:{x \<in> state. <k, h(f(x), g(x))> \<in> t} \<longmapsto>w {x \<in> state. <k, h(f1(x), g1(x))> \<in> t}"
 apply (rule_tac B = "{x \<in> state. <k, h (f1 (x), g (x))> \<in> t}" in LeadsTo_Trans)
 apply (blast intro: imp_LeadsTo_comp_right)
@@ -174,7 +174,7 @@
 
 lemma FollowsD:
 "F \<in> Follows(A, r, f, g)\<Longrightarrow>
-  F \<in> program & (\<exists>a. a \<in> A) & (\<forall>x \<in> state. f(x):A & g(x):A)"
+  F \<in> program \<and> (\<exists>a. a \<in> A) \<and> (\<forall>x \<in> state. f(x):A \<and> g(x):A)"
 apply (unfold Follows_def)
 apply (blast dest: IncreasingD)
 done
@@ -440,7 +440,7 @@
 lemma Always_munion:
 "\<lbrakk>F \<in> Always({s \<in> state. <f1(s),f(s)> \<in> MultLe(A,r)});
           F \<in> Always({s \<in> state. <g1(s), g(s)> \<in> MultLe(A,r)});
-  \<forall>x \<in> state. f1(x):Mult(A)&f(x):Mult(A) & g1(x):Mult(A) & g(x):Mult(A)\<rbrakk>
+  \<forall>x \<in> state. f1(x):Mult(A)\<and>f(x):Mult(A) \<and> g1(x):Mult(A) \<and> g(x):Mult(A)\<rbrakk>
       \<Longrightarrow> F \<in> Always({s \<in> state. <f1(s) +# g1(s), f(s) +# g(s)> \<in> MultLe(A,r)})"
 apply (rule_tac h = munion in imp_Always_comp2, simp_all)
 apply (blast intro: munion_mono, simp_all)
@@ -456,8 +456,8 @@
 
 lemma Follows_msetsum_UN:
 "\<And>f. \<lbrakk>\<forall>i \<in> I. F \<in> Follows(Mult(A), MultLe(A, r), f'(i), f(i));
-  \<forall>s. \<forall>i \<in> I. multiset(f'(i, s)) & mset_of(f'(i, s))<=A &
-                        multiset(f(i, s)) & mset_of(f(i, s))<=A ;
+  \<forall>s. \<forall>i \<in> I. multiset(f'(i, s)) \<and> mset_of(f'(i, s))<=A \<and>
+                        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),
--- a/src/ZF/UNITY/GenPrefix.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/GenPrefix.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -17,7 +17,7 @@
 
 definition (*really belongs in ZF/Trancl*)
   part_order :: "[i, i] => o"  where
-  "part_order(A, r) \<equiv> refl(A,r) & trans[A](r) & antisym(r)"
+  "part_order(A, r) \<equiv> refl(A,r) \<and> trans[A](r) \<and> antisym(r)"
 
 consts
   gen_prefix :: "[i, i] => i"
@@ -80,8 +80,8 @@
 lemma Cons_gen_prefix_aux:
   "\<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 & ys' = Cons(y,ys) &
-       <x,y>:r & <xs, ys> \<in> gen_prefix(A, r)))"
+       (\<exists>y ys. y \<in> A \<and> ys' = Cons(y,ys) \<and>
+       <x,y>:r \<and> <xs, ys> \<in> gen_prefix(A, r)))"
 apply (erule gen_prefix.induct)
 prefer 3 apply (force intro: gen_prefix.append, auto)
 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 & y \<in> A & <x,y>:r & <xs,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))"
 apply (auto intro: gen_prefix.prepend)
 done
 declare Cons_gen_prefix_Cons [iff]
@@ -213,7 +213,7 @@
 prefer 2 apply (blast dest: gen_prefix.dom_subset [THEN subsetD])
 apply (drule_tac psi = "<ys @ zs, xs> \<in> gen_prefix (A,r) " in asm_rl)
 apply simp
-apply (subgoal_tac "length (ys @ zs) = length (ys) #+ length (zs) &ys \<in> list (A) &xs \<in> list (A) ")
+apply (subgoal_tac "length (ys @ zs) = length (ys) #+ length (zs) \<and>ys \<in> list (A) \<and>xs \<in> list (A) ")
 prefer 2 apply (blast intro: length_app dest: gen_prefix.dom_subset [THEN subsetD])
 apply (drule gen_prefix_length_le)+
 apply clarify
@@ -240,7 +240,7 @@
 
 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) & z \<in> A & <z,y>:r & <zs,ys> \<in> gen_prefix(A,r)))"
+      (xs=[] | (\<exists>z zs. xs=Cons(z,zs) \<and> z \<in> A \<and> <z,y>:r \<and> <zs,ys> \<in> gen_prefix(A,r)))"
 apply (induct_tac "xs", auto)
 done
 
@@ -276,7 +276,7 @@
 (* Append case is hardest *)
 apply (frule gen_prefix_length_le [THEN le_iff [THEN iffD1]])
 apply (frule gen_prefix.dom_subset [THEN subsetD], clarify)
-apply (subgoal_tac "length (xs) :nat&length (ys) :nat &length (zs) :nat")
+apply (subgoal_tac "length (xs) :nat\<and>length (ys) :nat \<and>length (zs) :nat")
 prefer 2 apply (blast intro: length_type, clarify)
 apply (simp_all add: nth_append length_type length_app)
 apply (rule conjI)
@@ -325,7 +325,7 @@
 done
 
 lemma gen_prefix_iff_nth: "(<xs,ys> \<in> gen_prefix(A,r)) \<longleftrightarrow>
-      (xs \<in> list(A) & ys \<in> list(A) & length(xs) \<le> length(ys) &
+      (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)
 apply (frule gen_prefix.dom_subset [THEN subsetD])
@@ -367,7 +367,7 @@
 
 apply (unfold prefix_def)
 apply (erule gen_prefix.induct)
-apply (subgoal_tac [3] "xs \<in> list (A) &ys \<in> list (A) ")
+apply (subgoal_tac [3] "xs \<in> list (A) \<and>ys \<in> list (A) ")
 prefer 4 apply (blast dest: gen_prefix.dom_subset [THEN subsetD])
 apply (auto simp add: set_of_list_append)
 done
@@ -392,7 +392,7 @@
 declare prefix_Nil [iff]
 
 lemma Cons_prefix_Cons:
-"<Cons(x,xs), Cons(y,ys)> \<in> prefix(A) \<longleftrightarrow> (x=y & <xs,ys> \<in> prefix(A) & y \<in> A)"
+"<Cons(x,xs), Cons(y,ys)> \<in> prefix(A) \<longleftrightarrow> (x=y \<and> <xs,ys> \<in> prefix(A) \<and> y \<in> A)"
 apply (unfold prefix_def, auto)
 done
 declare Cons_prefix_Cons [iff]
@@ -422,7 +422,7 @@
 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) & <zs,ys> \<in> prefix(A)))"
+  (xs=[] | (\<exists>zs. xs=Cons(y,zs) \<and> <zs,ys> \<in> prefix(A)))"
 apply (unfold prefix_def)
 apply (auto simp add: gen_prefix_Cons)
 done
@@ -457,7 +457,7 @@
      "<xs,ys> \<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) & xs \<in> list(A)")
+apply (subgoal_tac [!] "ys \<in> list(A) \<and> xs \<in> list(A)")
 apply (auto dest: gen_prefix.dom_subset [THEN subsetD]
             simp add: length_type)
 apply (subgoal_tac "length (zs) =0")
@@ -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) & xs \<in> list(A)"
+    "<xs,zs> \<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) ")
@@ -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 & <us,zs> \<in> prefix(A)))"
+  (<xs,ys> \<in> prefix(A) | (\<exists>us. xs = ys@us \<and> <us,zs> \<in> prefix(A)))"
 apply (erule list_append_induct, force, clarify)
 apply (rule iffI)
 apply (simp add: add: app_assoc [symmetric])
@@ -646,7 +646,7 @@
 apply (erule natE, auto)
 done
 
-lemma prefix_take_iff: "<xs,ys> \<in> prefix(A) \<longleftrightarrow> (xs=take(length(xs), ys) & xs \<in> list(A) & ys \<in> list(A))"
+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))"
 apply (rule iffI)
 apply (frule prefix_type [THEN subsetD])
 apply (blast intro: prefix_imp_take, clarify)
--- a/src/ZF/UNITY/Guar.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/Guar.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -38,23 +38,23 @@
 
 definition
    ex_prop :: "i => o"  where
-   "ex_prop(X) \<equiv> X<=program &
+   "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(X) \<equiv> X<=program &
+  "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(X) \<equiv> X<=program &
-    (SKIP \<in> X & (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<in> X & G \<in> X \<longrightarrow> (F \<squnion> G) \<in> X))"
+   "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(X) \<equiv> X<=program &
-    (SKIP \<in> X & (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow>(F \<in> X & G \<in> X) \<longleftrightarrow> (F \<squnion> G \<in> X)))"
+   "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
@@ -69,7 +69,7 @@
 definition
   (* Weakest existential property stronger than X *)
    wx :: "i =>i"  where
-   "wx(X) \<equiv> \<Union>({Y \<in> Pow(program). Y<=X & ex_prop(Y)})"
+   "wx(X) \<equiv> \<Union>({Y \<in> Pow(program). Y<=X \<and> ex_prop(Y)})"
 
 definition
   (*Ill-defined programs can arise through "\<squnion>"*)
@@ -79,7 +79,7 @@
 definition
   refines :: "[i, i, i] => o" (\<open>(3_ refines _ wrt _)\<close> [10,10,10] 10)  where
   "G refines F wrt X \<equiv>
-   \<forall>H \<in> program. (F ok H  & G ok H & F \<squnion> H \<in> welldef \<inter> X)
+   \<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
@@ -111,11 +111,11 @@
 apply (auto intro: ok_sym)
 done
 
-(*Chandy & Sanders take this as a definition*)
+(*Chandy \<and> Sanders take this as a definition*)
 
 lemma ex_prop_finite:
-     "ex_prop(X) \<longleftrightarrow> (X\<subseteq>program &  
-  (\<forall>GG \<in> Fin(program). GG \<inter> X \<noteq> 0 & OK(GG,(%G. G))\<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X))"
+     "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))"
 apply auto
 apply (blast intro: ex1 ex2 dest: ex_imp_subset_program)+
 done
@@ -123,7 +123,7 @@
 (* Equivalent definition of ex_prop given at the end of section 3*)
 lemma ex_prop_equiv: 
 "ex_prop(X) \<longleftrightarrow>  
-  X\<subseteq>program & (\<forall>G \<in> program. (G \<in> X \<longleftrightarrow> (\<forall>H \<in> program. (G component_of H) \<longrightarrow> H \<in> X)))"
+  X\<subseteq>program \<and> (\<forall>G \<in> program. (G \<in> X \<longleftrightarrow> (\<forall>H \<in> program. (G component_of H) \<longrightarrow> H \<in> X)))"
 apply (unfold ex_prop_def component_of_def, safe, force, force, blast) 
 apply (subst Join_commute)
 apply (blast intro: ok_sym) 
@@ -138,7 +138,7 @@
 
 lemma uv1 [rule_format]: 
      "GG \<in> Fin(program) \<Longrightarrow>  
-      (uv_prop(X)\<longrightarrow> GG \<subseteq> X & OK(GG, (%G. G)) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X)"
+      (uv_prop(X)\<longrightarrow> GG \<subseteq> X \<and> OK(GG, (%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 & OK(GG,(%G. G)) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X)
+      (\<forall>GG \<in> Fin(program). GG \<subseteq> X \<and> OK(GG,(%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+) 
@@ -154,10 +154,10 @@
 apply (force dest: ok_sym simp add: OK_iff_ok) 
 done
 
-(*Chandy & Sanders take this as a definition*)
+(*Chandy \<and> Sanders take this as a definition*)
 lemma uv_prop_finite: 
-"uv_prop(X) \<longleftrightarrow> X\<subseteq>program &  
-  (\<forall>GG \<in> Fin(program). GG \<subseteq> X & OK(GG, %G. G) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in>  X)"
+"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)"
 apply auto
 apply (blast dest: uv_imp_subset_program)
 apply (blast intro: uv1)
@@ -388,7 +388,7 @@
 lemma refines_refl: "F refines F wrt X"
 by (unfold refines_def, blast)
 
-(* More results on guarantees, added by Sidi Ehmety from Chandy & Sander, section 6 *)
+(* More results on guarantees, added by Sidi Ehmety from Chandy \<and> Sander, section 6 *)
 
 lemma wg_type: "wg(F, X) \<subseteq> program"
 by (unfold wg_def, auto)
@@ -396,14 +396,14 @@
 lemma guarantees_type: "X guarantees Y \<subseteq> program"
 by (unfold guar_def, auto)
 
-lemma wgD2: "G \<in> wg(F, X) \<Longrightarrow> G \<in> program & F \<in> program"
+lemma wgD2: "G \<in> wg(F, X) \<Longrightarrow> G \<in> program \<and> F \<in> program"
 apply (unfold wg_def, auto)
 apply (blast dest: guarantees_type [THEN subsetD])
 done
 
 lemma guarantees_equiv: 
 "(F \<in> X guarantees Y) \<longleftrightarrow>  
-   F \<in> program & (\<forall>H \<in> program. H \<in> X \<longrightarrow> (F component_of H \<longrightarrow> H \<in> Y))"
+   F \<in> program \<and> (\<forall>H \<in> program. H \<in> X \<longrightarrow> (F component_of H \<longrightarrow> H \<in> Y))"
 by (unfold guar_def component_of_def, force) 
 
 lemma wg_weakest:
@@ -415,7 +415,7 @@
 
 lemma wg_equiv:
      "H \<in> wg(F,X) \<longleftrightarrow> 
-      ((F component_of H \<longrightarrow> H \<in> X) & F \<in> program & H \<in> program)"
+      ((F component_of H \<longrightarrow> H \<in> X) \<and> F \<in> program \<and> H \<in> program)"
 apply (simp add: wg_def guarantees_equiv)
 apply (rule iffI, safe)
 apply (rule_tac [4] x = "{H}" in bexI)
@@ -423,7 +423,7 @@
 done
 
 lemma component_of_wg: 
-    "F component_of H \<Longrightarrow> H \<in> wg(F,X) \<longleftrightarrow> (H \<in> X & F \<in> program & H \<in> program)"
+    "F component_of H \<Longrightarrow> H \<in> wg(F,X) \<longleftrightarrow> (H \<in> X \<and> F \<in> program \<and> H \<in> program)"
 by (simp (no_asm_simp) add: wg_equiv)
 
 lemma wg_finite [rule_format]: 
@@ -439,7 +439,7 @@
 done
 
 lemma wg_ex_prop:
-     "ex_prop(X) \<Longrightarrow> (F \<in> X) \<longleftrightarrow> (\<forall>H \<in> program. H \<in> wg(F,X) & F \<in> program)"
+     "ex_prop(X) \<Longrightarrow> (F \<in> X) \<longleftrightarrow> (\<forall>H \<in> program. H \<in> wg(F,X) \<and> F \<in> program)"
 apply (simp (no_asm_use) add: ex_prop_equiv wg_equiv)
 apply blast
 done
--- a/src/ZF/UNITY/Increasing.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/Increasing.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -13,13 +13,13 @@
 definition
   increasing :: "[i, i, i=>i] => 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})) &
+    {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[A](r, f) \<equiv>
-    {F \<in> program. (\<forall>k \<in> A. F \<in> Stable({s \<in> state. <k, f(s)> \<in> r})) &
+    {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)
@@ -40,14 +40,14 @@
 by (unfold increasing_def, blast)
 
 lemma increasingD: 
-"F \<in> increasing[A](r,f) \<Longrightarrow> F \<in> program & (\<exists>a. a \<in> A) & (\<forall>s \<in> state. f(s):A)"
+"F \<in> increasing[A](r,f) \<Longrightarrow> F \<in> program \<and> (\<exists>a. a \<in> A) \<and> (\<forall>s \<in> state. f(s):A)"
 apply (unfold increasing_def)
 apply (subgoal_tac "\<exists>x. x \<in> state")
 apply (auto dest: stable_type [THEN subsetD] intro: st0_in_state)
 done
 
 lemma increasing_constant [simp]: 
- "F \<in> increasing[A](r, %s. c) \<longleftrightarrow> F \<in> program & c \<in> A"
+ "F \<in> increasing[A](r, %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)
@@ -109,14 +109,14 @@
 by (unfold Increasing_def, blast)
 
 lemma IncreasingD: 
-"F \<in> Increasing[A](r, f) \<Longrightarrow> F \<in> program & (\<exists>a. a \<in> A) & (\<forall>s \<in> state. f(s):A)"
+"F \<in> Increasing[A](r, f) \<Longrightarrow> F \<in> program \<and> (\<exists>a. a \<in> A) \<and> (\<forall>s \<in> state. f(s):A)"
 apply (unfold Increasing_def)
 apply (subgoal_tac "\<exists>x. x \<in> state")
 apply (auto intro: st0_in_state)
 done
 
 lemma Increasing_constant [simp]: 
-     "F \<in> Increasing[A](r, %s. c) \<longleftrightarrow> F \<in> program & (c \<in> A)"
+     "F \<in> Increasing[A](r, %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
@@ -127,7 +127,7 @@
 apply (unfold Increasing_def Stable_def Constrains_def part_order_def 
        constrains_def mono1_def metacomp_def, safe)
 apply (simp_all add: ActsD)
-apply (subgoal_tac "xb \<in> state & xa \<in> state")
+apply (subgoal_tac "xb \<in> state \<and> xa \<in> state")
  prefer 2 apply (simp add: ActsD)
 apply (subgoal_tac "<f (xb), f (xb) >:r")
 prefer 2 apply (force simp add: refl_def)
@@ -168,9 +168,9 @@
        part_order_def constrains_def mono2_def, clarify, simp)
 apply clarify
 apply (rename_tac xa xb)
-apply (subgoal_tac "xb \<in> state & xa \<in> state")
+apply (subgoal_tac "xb \<in> state \<and> xa \<in> state")
  prefer 2 apply (blast dest!: ActsD)
-apply (subgoal_tac "<f (xb), f (xb) >:r & <g (xb), g (xb) >:s")
+apply (subgoal_tac "<f (xb), f (xb) >:r \<and> <g (xb), g (xb) >:s")
 prefer 2 apply (force simp add: refl_def)
 apply (rotate_tac 6)
 apply (drule_tac x = "f (xb) " in bspec)
@@ -199,9 +199,9 @@
 apply (unfold Increasing_def stable_def 
        part_order_def constrains_def mono2_def Stable_def Constrains_def, safe)
 apply (simp_all add: ActsD)
-apply (subgoal_tac "xa \<in> state & x \<in> state")
+apply (subgoal_tac "xa \<in> state \<and> x \<in> state")
 prefer 2 apply (blast dest!: ActsD)
-apply (subgoal_tac "<f (xa), f (xa) >:r & <g (xa), g (xa) >:s")
+apply (subgoal_tac "<f (xa), f (xa) >:r \<and> <g (xa), g (xa) >:s")
 prefer 2 apply (force simp add: refl_def)
 apply (rotate_tac 6)
 apply (drule_tac x = "f (xa) " in bspec)
--- a/src/ZF/UNITY/Merge.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/Merge.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -38,7 +38,7 @@
      (\<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) &
+        (%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))"
 
@@ -70,16 +70,16 @@
     and A    \<comment> \<open>the type of items being merged\<close>
     and M
  assumes var_assumes [simp]:
-           "(\<forall>n. In(n):var) & Out \<in> var & iOut \<in> var"
+           "(\<forall>n. In(n):var) \<and> Out \<in> var \<and> iOut \<in> var"
      and all_distinct_vars:
            "\<forall>n. all_distinct([In(n), Out, iOut])"
      and type_assumes [simp]:
-           "(\<forall>n. type_of(In(n))=list(A)) &
-            type_of(Out)=list(A) &
+           "(\<forall>n. type_of(In(n))=list(A)) \<and>
+            type_of(Out)=list(A) \<and>
             type_of(iOut)=list(nat)"
      and default_val_assumes [simp]: 
-           "(\<forall>n. default_val(In(n))=Nil) &
-            default_val(Out)=Nil &
+           "(\<forall>n. default_val(In(n))=Nil) \<and>
+            default_val(Out)=Nil \<and>
             default_val(iOut)=Nil"
      and merge_spec:  "M \<in> merge_spec(A, In, Out, iOut)"
 
@@ -114,8 +114,8 @@
 
 lemma (in merge) M_ok_iff: 
      "G \<in> program \<Longrightarrow>  
-       M ok G \<longleftrightarrow> (G \<in> preserves(lift(Out)) &   
-                   G \<in> preserves(lift(iOut)) & M \<in> Allowed(G))"
+       M ok G \<longleftrightarrow> (G \<in> preserves(lift(Out)) \<and>   
+                   G \<in> preserves(lift(iOut)) \<and> M \<in> Allowed(G))"
 apply (cut_tac merge_spec)
 apply (auto simp add: merge_Allowed ok_iff_Allowed)
 done
@@ -147,14 +147,14 @@
     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,  
-      {k \<in> nat. k < length(s`iOut) & nth(k, s`iOut)=i})),  
+      {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]) 
 apply (rule_tac [2] state_AlwaysI [THEN Always_weaken]) 
 apply (rule Always_Int_I [OF merge_Always_Out_eq_iOut merge_Bounded], auto)
 apply (subst bag_of_sublist_UN_disjoint [symmetric])
 apply (auto simp add: nat_into_Finite set_of_list_conv_nth  [OF iOut_value_type])
-apply (subgoal_tac " (\<Union>i \<in> Nclients. {k \<in> nat. k < length (x`iOut) & nth (k, x`iOut) = i}) = length (x`iOut) ")
+apply (subgoal_tac " (\<Union>i \<in> Nclients. {k \<in> nat. k < length (x`iOut) \<and> nth (k, x`iOut) = i}) = length (x`iOut) ")
 apply (auto simp add: sublist_upt_eq_take [OF Out_value_type] 
                       length_type  [OF iOut_value_type]  
                       take_all [OF _ Out_value_type] 
--- a/src/ZF/UNITY/Monotonicity.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/Monotonicity.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -14,7 +14,7 @@
 definition
   mono1 :: "[i, i, i, i, i=>i] => 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) & (\<forall>x \<in> A. f(x) \<in> B)"
+    (\<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)"
 
   (* monotonicity of a 2-place meta-function f *)
 
@@ -22,7 +22,7 @@
   mono2 :: "[i, i, i, i, i, i, [i,i]=>i] => 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 & <u,v> \<in> s \<longrightarrow> <f(x,u), f(y,v)> \<in> t) &
+              <x,y> \<in> r \<and> <u,v> \<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 *)
--- a/src/ZF/UNITY/MultisetSum.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/MultisetSum.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -11,7 +11,7 @@
 definition
   lcomm :: "[i, i, [i,i]=>i]=>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)))  &
+   (\<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
@@ -57,7 +57,7 @@
 (** msetsum **)
 
 lemma multiset_general_setsum: 
-     "\<lbrakk>C \<in> Fin(A); \<forall>x \<in> A. multiset(g(x))& mset_of(g(x))\<subseteq>B\<rbrakk>   
+     "\<lbrakk>C \<in> Fin(A); \<forall>x \<in> A. multiset(g(x))\<and> mset_of(g(x))\<subseteq>B\<rbrakk>   
       \<Longrightarrow> multiset(general_setsum(C, B -||> nat - {0}, 0, \<lambda>u v. u +# v, g))"
 apply (erule Fin_induct, auto)
 apply (subst general_setsum_cons)
@@ -68,7 +68,7 @@
 by (simp add: msetsum_def)
 
 lemma msetsum_cons [simp]: 
-  "\<lbrakk>C \<in> Fin(A); a\<notin>C; a \<in> A; \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B\<rbrakk>   
+  "\<lbrakk>C \<in> Fin(A); a\<notin>C; a \<in> A; \<forall>x \<in> A. multiset(g(x)) \<and> mset_of(g(x))\<subseteq>B\<rbrakk>   
    \<Longrightarrow> msetsum(g, cons(a, C), B) = g(a) +#  msetsum(g, C, B)"
 apply (simp add: msetsum_def)
 apply (subst general_setsum_cons)
@@ -81,20 +81,20 @@
 by (simp add: msetsum_def)
 
 lemma mset_of_msetsum: 
-     "\<lbrakk>C \<in> Fin(A); \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B\<rbrakk>   
+     "\<lbrakk>C \<in> Fin(A); \<forall>x \<in> A. multiset(g(x)) \<and> mset_of(g(x))\<subseteq>B\<rbrakk>   
       \<Longrightarrow> mset_of(msetsum(g, C, B))\<subseteq>B"
 by (erule Fin_induct, auto)
 
 
 (*The reversed orientation looks more natural, but LOOPS as a simprule!*)
 lemma msetsum_Un_Int: 
-     "\<lbrakk>C \<in> Fin(A); D \<in> Fin(A); \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B\<rbrakk>
+     "\<lbrakk>C \<in> Fin(A); D \<in> Fin(A); \<forall>x \<in> A. multiset(g(x)) \<and> mset_of(g(x))\<subseteq>B\<rbrakk>
       \<Longrightarrow> msetsum(g, C \<union> D, B) +# msetsum(g, C \<inter> D, B)  
         = msetsum(g, C, B) +# msetsum(g, D, B)"
 apply (erule Fin_induct)
 apply (subgoal_tac [2] "cons (x, y) \<union> D = cons (x, y \<union> D) ")
 apply (auto simp add: msetsum_multiset)
-apply (subgoal_tac "y \<union> D \<in> Fin (A) & y \<inter> D \<in> Fin (A) ")
+apply (subgoal_tac "y \<union> D \<in> Fin (A) \<and> y \<inter> D \<in> Fin (A) ")
 apply clarify
 apply (case_tac "x \<in> D")
 apply (subgoal_tac [2] "cons (x, y) \<inter> D = y \<inter> D")
@@ -107,7 +107,7 @@
 
 lemma msetsum_Un_disjoint:
      "\<lbrakk>C \<in> Fin(A); D \<in> Fin(A); C \<inter> D = 0;  
-         \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B\<rbrakk>  
+         \<forall>x \<in> A. multiset(g(x)) \<and> mset_of(g(x))\<subseteq>B\<rbrakk>  
       \<Longrightarrow> msetsum(g, C \<union> D, B) = msetsum(g, C, B) +# msetsum(g,D, B)"
 apply (subst msetsum_Un_Int [symmetric])
 apply (auto simp add: msetsum_multiset)
@@ -119,7 +119,7 @@
  
 lemma msetsum_UN_disjoint [rule_format (no_asm)]:
      "\<lbrakk>I \<in> Fin(K); \<forall>i \<in> K. C(i) \<in> Fin(A)\<rbrakk> \<Longrightarrow>  
-      (\<forall>x \<in> A. multiset(f(x)) & mset_of(f(x))\<subseteq>B) \<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)"
 apply (erule Fin_induct, auto)
@@ -127,10 +127,10 @@
  prefer 2 apply blast
 apply (subgoal_tac "C(x) \<inter> (\<Union>i \<in> y. C (i)) = 0")
  prefer 2 apply blast
-apply (subgoal_tac " (\<Union>i \<in> y. C (i)):Fin (A) & C(x) :Fin (A) ")
+apply (subgoal_tac " (\<Union>i \<in> y. C (i)):Fin (A) \<and> C(x) :Fin (A) ")
 prefer 2 apply (blast intro: UN_Fin_lemma dest: FinD, clarify)
 apply (simp (no_asm_simp) add: msetsum_Un_disjoint)
-apply (subgoal_tac "\<forall>x \<in> K. multiset (msetsum (f, C(x), B)) & mset_of (msetsum (f, C(x), B)) \<subseteq> B")
+apply (subgoal_tac "\<forall>x \<in> K. multiset (msetsum (f, C(x), B)) \<and> mset_of (msetsum (f, C(x), B)) \<subseteq> B")
 apply (simp (no_asm_simp))
 apply clarify
 apply (drule_tac x = xa in bspec)
@@ -139,10 +139,10 @@
 
 lemma msetsum_addf: 
     "\<lbrakk>C \<in> Fin(A);  
-      \<forall>x \<in> A. multiset(f(x)) & mset_of(f(x))\<subseteq>B;  
-      \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B\<rbrakk> \<Longrightarrow> 
+      \<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)"
-apply (subgoal_tac "\<forall>x \<in> A. multiset (f(x) +# g(x)) & mset_of (f(x) +# g(x)) \<subseteq> 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) 
 done
@@ -156,10 +156,10 @@
 by (simp add: multiset_equality)
 
 lemma msetsum_Un: "\<lbrakk>C \<in> Fin(A); D \<in> Fin(A);  
-  \<forall>x \<in> A. multiset(f(x)) & mset_of(f(x)) \<subseteq> B\<rbrakk>  
+  \<forall>x \<in> A. multiset(f(x)) \<and> mset_of(f(x)) \<subseteq> B\<rbrakk>  
    \<Longrightarrow> msetsum(f, C \<union> D, B) =  
           msetsum(f, C, B) +# msetsum(f, D, B) -# msetsum(f, C \<inter> D, B)"
-apply (subgoal_tac "C \<union> D \<in> Fin (A) & C \<inter> D \<in> Fin (A) ")
+apply (subgoal_tac "C \<union> D \<in> Fin (A) \<and> C \<inter> D \<in> Fin (A) ")
 apply clarify
 apply (subst msetsum_Un_Int [symmetric])
 apply (simp_all (no_asm_simp) add: msetsum_multiset multiset_union_diff)
--- a/src/ZF/UNITY/Mutex.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/Mutex.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -28,65 +28,65 @@
 abbreviation "v \<equiv> Var([1,0])"
 
 axiomatization where \<comment> \<open>Type declarations\<close>
-  p_type:  "type_of(p)=bool & default_val(p)=0" and
-  m_type:  "type_of(m)=int  & default_val(m)=#0" and
-  n_type:  "type_of(n)=int  & default_val(n)=#0" and
-  u_type:  "type_of(u)=bool & default_val(u)=0" and
-  v_type:  "type_of(v)=bool & default_val(v)=0"
+  p_type:  "type_of(p)=bool \<and> default_val(p)=0" and
+  m_type:  "type_of(m)=int  \<and> default_val(m)=#0" and
+  n_type:  "type_of(n)=int  \<and> default_val(n)=#0" and
+  u_type:  "type_of(u)=bool \<and> default_val(u)=0" and
+  v_type:  "type_of(v)=bool \<and> default_val(v)=0"
 
 definition
   (** The program for process U **)
-  "U0 \<equiv> {<s,t>:state*state. t = s(u:=1, m:=#1) & s`m = #0}"
+  "U0 \<equiv> {<s,t>: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) &  s`m = #1}"
+  "U1 \<equiv> {<s,t>:state*state. t = s(p:= s`v, m:=#2) \<and>  s`m = #1}"
 
 definition
-  "U2 \<equiv> {<s,t>:state*state. t = s(m:=#3) & s`p=0 & s`m = #2}"
+  "U2 \<equiv> {<s,t>: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) & s`m = #3}"
+  "U3 \<equiv> {<s,t>: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) & s`m = #4}"
+  "U4 \<equiv> {<s,t>: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) & s`n = #0}"
+  "V0 \<equiv> {<s,t>: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) & s`n = #1}"
+  "V1 \<equiv> {<s,t>: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) & s`p=1 & s`n = #2}"
+  "V2 \<equiv> {<s,t>: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) & s`n = #3}"
+  "V3 \<equiv> {<s,t>: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) & s`n = #4}"
+  "V4 \<equiv> {<s,t>:state*state. t  = s (p:=0, n:=#0) \<and> s`n = #4}"
 
 definition
-  "Mutex \<equiv> mk_program({s:state. s`u=0 & s`v=0 & s`m = #0 & s`n = #0},
+  "Mutex \<equiv> mk_program({s:state. s`u=0 \<and> s`v=0 \<and> s`m = #0 \<and> s`n = #0},
               {U0, U1, U2, U3, U4, V0, V1, V2, V3, V4}, Pow(state*state))"
 
   (** The correct invariants **)
 
 definition
-  "IU \<equiv> {s:state. (s`u = 1\<longleftrightarrow>(#1 $\<le> s`m & s`m $\<le> #3))
-                     & (s`m = #3 \<longrightarrow> s`p=0)}"
+  "IU \<equiv> {s:state. (s`u = 1\<longleftrightarrow>(#1 $\<le> s`m \<and> s`m $\<le> #3))
+                     \<and> (s`m = #3 \<longrightarrow> s`p=0)}"
 
 definition
-  "IV \<equiv> {s:state. (s`v = 1 \<longleftrightarrow> (#1 $\<le> s`n & s`n $\<le> #3))
-                      & (s`n = #3 \<longrightarrow> s`p=1)}"
+  "IV \<equiv> {s:state. (s`v = 1 \<longleftrightarrow> (#1 $\<le> s`n \<and> s`n $\<le> #3))
+                      \<and> (s`n = #3 \<longrightarrow> s`p=1)}"
 
   (** The faulty invariant (for U alone) **)
 
 definition
-  "bad_IU \<equiv> {s:state. (s`u = 1 \<longleftrightarrow> (#1 $\<le> s`m & s`m  $\<le> #3))&
-                   (#3 $\<le> s`m & s`m $\<le> #4 \<longrightarrow> s`p=0)}"
+  "bad_IU \<equiv> {s:state. (s`u = 1 \<longleftrightarrow> (#1 $\<le> s`m \<and> s`m  $\<le> #3))\<and>
+                   (#3 $\<le> s`m \<and> s`m $\<le> #4 \<longrightarrow> s`p=0)}"
 
 
 (** Variables' types **)
@@ -174,7 +174,7 @@
 done
 
 (*The safety property: mutual exclusion*)
-lemma mutual_exclusion: "Mutex \<in> Always({s \<in> state. \<not>(s`m = #3 & s`n = #3)})"
+lemma mutual_exclusion: "Mutex \<in> Always({s \<in> state. \<not>(s`m = #3 \<and> s`n = #3)})"
 apply (rule Always_weaken)
 apply (rule Always_Int_I [OF IU IV], auto)
 done
@@ -206,10 +206,10 @@
 by (unfold op_Unless_def Mutex_def, safety)
 
 lemma U_F1:
-     "Mutex \<in> {s \<in> state. s`m=#1} \<longmapsto>w {s \<in> state. s`p = s`v & s`m = #2}"
+     "Mutex \<in> {s \<in> state. s`m=#1} \<longmapsto>w {s \<in> state. s`p = s`v \<and> s`m = #2}"
 by (unfold Mutex_def, ensures U1)
 
-lemma U_F2: "Mutex \<in> {s \<in> state. s`p =0 & s`m = #2} \<longmapsto>w {s \<in> state. s`m = #3}"
+lemma U_F2: "Mutex \<in> {s \<in> state. s`p =0 \<and> s`m = #2} \<longmapsto>w {s \<in> state. s`m = #3}"
 apply (cut_tac IU)
 apply (unfold Mutex_def, ensures U2)
 done
@@ -232,7 +232,7 @@
 lemma U_lemma1: "Mutex \<in> {s \<in> state. s`m = #1} \<longmapsto>w {s \<in> state. s`p =1}"
 by (rule LeadsTo_Trans [OF U_F1 [THEN LeadsTo_weaken_R] U_lemma2], blast)
 
-lemma eq_123: "i \<in> int \<Longrightarrow> (#1 $\<le> i & i $\<le> #3) \<longleftrightarrow> (i=#1 | i=#2 | i=#3)"
+lemma eq_123: "i \<in> int \<Longrightarrow> (#1 $\<le> i \<and> i $\<le> #3) \<longleftrightarrow> (i=#1 | i=#2 | i=#3)"
 apply auto
 apply (auto simp add: neq_iff_zless)
 apply (drule_tac [4] j = "#3" and i = i in zle_zless_trans)
@@ -243,7 +243,7 @@
 done
 
 
-lemma U_lemma123: "Mutex \<in> {s \<in> state. #1 $\<le> s`m & s`m $\<le> #3} \<longmapsto>w {s \<in> state. s`p=1}"
+lemma U_lemma123: "Mutex \<in> {s \<in> state. #1 $\<le> s`m \<and> s`m $\<le> #3} \<longmapsto>w {s \<in> state. s`p=1}"
 by (simp add: eq_123 Collect_disj_eq LeadsTo_Un_distrib U_lemma1 U_lemma2 U_F3)
 
 
@@ -257,10 +257,10 @@
 lemma V_F0: "Mutex \<in> {s \<in> state. s`n=#2} Unless {s \<in> state. s`n=#3}"
 by (unfold op_Unless_def Mutex_def, safety)
 
-lemma V_F1: "Mutex \<in> {s \<in> state. s`n=#1} \<longmapsto>w {s \<in> state. s`p = not(s`u) & s`n = #2}"
+lemma V_F1: "Mutex \<in> {s \<in> state. s`n=#1} \<longmapsto>w {s \<in> state. s`p = not(s`u) \<and> s`n = #2}"
 by (unfold Mutex_def, ensures "V1")
 
-lemma V_F2: "Mutex \<in> {s \<in> state. s`p=1 & s`n = #2} \<longmapsto>w {s \<in> state. s`n = #3}"
+lemma V_F2: "Mutex \<in> {s \<in> state. s`p=1 \<and> s`n = #2} \<longmapsto>w {s \<in> state. s`n = #3}"
 apply (cut_tac IV)
 apply (unfold Mutex_def, ensures "V2")
 done
@@ -282,7 +282,7 @@
 lemma V_lemma1: "Mutex \<in> {s \<in> state. s`n = #1} \<longmapsto>w {s \<in> state. s`p = 0}"
 by (rule LeadsTo_Trans [OF V_F1 [THEN LeadsTo_weaken_R] V_lemma2], blast)
 
-lemma V_lemma123: "Mutex \<in> {s \<in> state. #1 $\<le> s`n & s`n $\<le> #3} \<longmapsto>w {s \<in> state. s`p = 0}"
+lemma V_lemma123: "Mutex \<in> {s \<in> state. #1 $\<le> s`n \<and> s`n $\<le> #3} \<longmapsto>w {s \<in> state. s`p = 0}"
 by (simp add: eq_123 Collect_disj_eq LeadsTo_Un_distrib V_lemma1 V_lemma2 V_F3)
 
 (*Misra's F4*)
--- a/src/ZF/UNITY/State.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/State.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -51,7 +51,7 @@
 
 (* Union *)
 
-lemma st_set_Un_iff [iff]: "st_set(A \<union> B) \<longleftrightarrow> st_set(A) & st_set(B)"
+lemma st_set_Un_iff [iff]: "st_set(A \<union> B) \<longleftrightarrow> st_set(A) \<and> st_set(B)"
 by (simp add: st_set_def, auto)
 
 lemma st_set_Union_iff [iff]: "st_set(\<Union>(S)) \<longleftrightarrow> (\<forall>A \<in> S. st_set(A))"
@@ -99,7 +99,7 @@
 lemma st_set_compl [simp]: "st_set(st_compl(A))"
 by (simp add: st_compl_def, auto)
 
-lemma st_compl_iff [simp]: "x \<in> st_compl(A) \<longleftrightarrow> x \<in> state & x \<notin> A"
+lemma st_compl_iff [simp]: "x \<in> st_compl(A) \<longleftrightarrow> x \<in> state \<and> x \<notin> A"
 by (simp add: st_compl_def)
 
 lemma st_compl_Collect [simp]:
@@ -108,7 +108,7 @@
 
 (*For using "disjunction" (union over an index set) to eliminate a variable.*)
 lemma UN_conj_eq:
-     "\<forall>d\<in>D. f(d) \<in> A \<Longrightarrow> (\<Union>k\<in>A. {d\<in>D. P(d) & f(d) = k}) = {d\<in>D. P(d)}"
+     "\<forall>d\<in>D. f(d) \<in> A \<Longrightarrow> (\<Union>k\<in>A. {d\<in>D. P(d) \<and> f(d) = k}) = {d\<in>D. P(d)}"
 by blast
 
 end
--- a/src/ZF/UNITY/SubstAx.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/SubstAx.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -153,14 +153,14 @@
 done
 
 (** Distributive laws **)
-lemma LeadsTo_Un_distrib: "(F \<in> (A \<union> B) \<longmapsto>w C)  \<longleftrightarrow> (F \<in> A \<longmapsto>w C & F \<in> B \<longmapsto>w C)"
+lemma LeadsTo_Un_distrib: "(F \<in> (A \<union> B) \<longmapsto>w C)  \<longleftrightarrow> (F \<in> A \<longmapsto>w C \<and> F \<in> B \<longmapsto>w C)"
 by (blast intro: LeadsTo_Un LeadsTo_weaken_L)
 
-lemma LeadsTo_UN_distrib: "(F \<in> (\<Union>i \<in> I. A(i)) \<longmapsto>w B) \<longleftrightarrow>  (\<forall>i \<in> I. F \<in> A(i) \<longmapsto>w B) & F \<in> program"
+lemma LeadsTo_UN_distrib: "(F \<in> (\<Union>i \<in> I. A(i)) \<longmapsto>w B) \<longleftrightarrow>  (\<forall>i \<in> I. F \<in> A(i) \<longmapsto>w B) \<and> F \<in> program"
 by (blast dest: LeadsTo_type [THEN subsetD]
              intro: LeadsTo_UN LeadsTo_weaken_L)
 
-lemma LeadsTo_Union_distrib: "(F \<in> \<Union>(S) \<longmapsto>w B)  \<longleftrightarrow>  (\<forall>A \<in> S. F \<in> A \<longmapsto>w B) & F \<in> program"
+lemma LeadsTo_Union_distrib: "(F \<in> \<Union>(S) \<longmapsto>w B)  \<longleftrightarrow>  (\<forall>A \<in> S. F \<in> A \<longmapsto>w B) \<and> F \<in> program"
 by (blast dest: LeadsTo_type [THEN subsetD]
              intro: LeadsTo_Union LeadsTo_weaken_L)
 
@@ -353,7 +353,7 @@
                 ORELSE   (*subgoal may involve LeadsTo, leadsTo or ensures*)
                 REPEAT (ares_tac ctxt [@{thm LeadsTo_Basis}, @{thm leadsTo_Basis},
                                   @{thm EnsuresI}, @{thm ensuresI}] 1),
-            (*now there are two subgoals: co & transient*)
+            (*now there are two subgoals: co \<and> transient*)
             simp_tac (ctxt addsimps (Named_Theorems.get ctxt \<^named_theorems>\<open>program\<close>)) 2,
             Rule_Insts.res_inst_tac ctxt
               [((("act", 0), Position.none), sact)] [] @{thm transientI} 2,
--- a/src/ZF/UNITY/UNITY.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/UNITY.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -16,7 +16,7 @@
   program  :: i  where
   "program \<equiv> {<init, acts, allowed>:
                Pow(state) * Pow(Pow(state*state)) * Pow(Pow(state*state)).
-               id(state) \<in> acts & id(state) \<in> allowed}"
+               id(state) \<in> acts \<and> id(state) \<in> allowed}"
 
 definition
   mk_program :: "[i,i,i]=>i"  where
@@ -69,7 +69,7 @@
   "initially(A) \<equiv> {F \<in> program. Init(F)\<subseteq>A}"
 
 definition "constrains" :: "[i, i] => i"  (infixl \<open>co\<close> 60)  where
-  "A co B \<equiv> {F \<in> program. (\<forall>act \<in> Acts(F). act``A\<subseteq>B) & st_set(A)}"
+  "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>
 
@@ -176,11 +176,11 @@
 by (simp add: RawAllowedActs_type AllowedActs_def)
 
 text\<open>Needed in Behaviors\<close>
-lemma ActsD: "\<lbrakk>act \<in> Acts(F); <s,s'> \<in> act\<rbrakk> \<Longrightarrow> s \<in> state & s' \<in> state"
+lemma ActsD: "\<lbrakk>act \<in> Acts(F); <s,s'> \<in> act\<rbrakk> \<Longrightarrow> s \<in> state \<and> s' \<in> state"
 by (blast dest: Acts_type [THEN subsetD])
 
 lemma AllowedActsD:
-     "\<lbrakk>act \<in> AllowedActs(F); <s,s'> \<in> act\<rbrakk> \<Longrightarrow> s \<in> state & s' \<in> state"
+     "\<lbrakk>act \<in> AllowedActs(F); <s,s'> \<in> act\<rbrakk> \<Longrightarrow> s \<in> state \<and> s' \<in> state"
 by (blast dest: AllowedActs_type [THEN subsetD])
 
 subsection\<open>Simplification rules involving \<^term>\<open>state\<close>, \<^term>\<open>Init\<close>, 
@@ -307,7 +307,7 @@
 
 lemma program_equality_iff:
     "\<lbrakk>F \<in> program; G \<in> program\<rbrakk> \<Longrightarrow>(F=G)  \<longleftrightarrow>
-     (Init(F) = Init(G) & Acts(F) = Acts(G) & AllowedActs(F) = AllowedActs(G))"
+     (Init(F) = Init(G) \<and> Acts(F) = Acts(G) \<and> AllowedActs(F) = AllowedActs(G))"
 by (blast intro: program_equalityI program_equalityE)
 
 subsection\<open>These rules allow "lazy" definition expansion\<close>
@@ -328,8 +328,8 @@
 
 lemma def_prg_simps:
     "\<lbrakk>F \<equiv> mk_program (init,acts,allowed)\<rbrakk>
-     \<Longrightarrow> Init(F) = init \<inter> state & 
-         Acts(F) = cons(id(state), acts \<inter> Pow(state*state)) &
+     \<Longrightarrow> Init(F) = init \<inter> state \<and> 
+         Acts(F) = cons(id(state), acts \<inter> Pow(state*state)) \<and>
          AllowedActs(F) = cons(id(state), allowed \<inter> Pow(state*state))"
 by auto
 
@@ -337,7 +337,7 @@
 text\<open>An action is expanded only if a pair of states is being tested against it\<close>
 lemma def_act_simp:
      "\<lbrakk>act \<equiv> {<s,s'> \<in> A*B. P(s, s')}\<rbrakk>
-      \<Longrightarrow> (<s,s'> \<in> act) \<longleftrightarrow> (<s,s'> \<in> A*B & P(s, s'))"
+      \<Longrightarrow> (<s,s'> \<in> act) \<longleftrightarrow> (<s,s'> \<in> A*B \<and> P(s, s'))"
 by auto
 
 text\<open>A set is expanded only if an element is being tested against it\<close>
@@ -359,21 +359,21 @@
    "F \<in> A co B \<Longrightarrow> \<forall>act \<in> Acts(F). act``A\<subseteq>B"
 by (force simp add: constrains_def)
 
-lemma constrainsD2: "F \<in> A co B \<Longrightarrow> F \<in> program & st_set(A)"
+lemma constrainsD2: "F \<in> A co B \<Longrightarrow> F \<in> program \<and> st_set(A)"
 by (force simp add: constrains_def)
 
 lemma constrains_empty [iff]: "F \<in> 0 co B \<longleftrightarrow> F \<in> program"
 by (force simp add: constrains_def st_set_def)
 
-lemma constrains_empty2 [iff]: "(F \<in> A co 0) \<longleftrightarrow> (A=0 & F \<in> program)"
+lemma constrains_empty2 [iff]: "(F \<in> A co 0) \<longleftrightarrow> (A=0 \<and> F \<in> program)"
 by (force simp add: constrains_def st_set_def)
 
-lemma constrains_state [iff]: "(F \<in> state co B) \<longleftrightarrow> (state\<subseteq>B & F \<in> program)"
+lemma constrains_state [iff]: "(F \<in> state co B) \<longleftrightarrow> (state\<subseteq>B \<and> F \<in> program)"
 apply (cut_tac F = F in Acts_type)
 apply (force simp add: constrains_def st_set_def)
 done
 
-lemma constrains_state2 [iff]: "F \<in> A co state \<longleftrightarrow> (F \<in> program & st_set(A))"
+lemma constrains_state2 [iff]: "F \<in> A co state \<longleftrightarrow> (F \<in> program \<and> st_set(A))"
 apply (cut_tac F = F in Acts_type)
 apply (force simp add: constrains_def st_set_def)
 done
@@ -500,7 +500,7 @@
 lemma stableD: "F \<in> stable(A) \<Longrightarrow> F \<in> A co A"
 by (unfold stable_def, assumption)
 
-lemma stableD2: "F \<in> stable(A) \<Longrightarrow> F \<in> program & st_set(A)"
+lemma stableD2: "F \<in> stable(A) \<Longrightarrow> F \<in> program \<and> st_set(A)"
 by (unfold stable_def constrains_def, auto)
 
 lemma stable_state [simp]: "stable(state) = program"
@@ -571,10 +571,10 @@
 apply (frule stable_type [THEN subsetD], auto)
 done
 
-lemma invariantD: "F \<in> invariant(A) \<Longrightarrow> Init(F)\<subseteq>A & F \<in> stable(A)"
+lemma invariantD: "F \<in> invariant(A) \<Longrightarrow> Init(F)\<subseteq>A \<and> F \<in> stable(A)"
 by (unfold invariant_def initially_def, auto)
 
-lemma invariantD2: "F \<in> invariant(A) \<Longrightarrow> F \<in> program & st_set(A)"
+lemma invariantD2: "F \<in> invariant(A) \<Longrightarrow> F \<in> program \<and> st_set(A)"
 apply (unfold invariant_def)
 apply (blast dest: stableD2)
 done
--- a/src/ZF/UNITY/Union.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/Union.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -16,7 +16,7 @@
 definition
   (*FIXME: conjoin Init(F) \<inter> Init(G) \<noteq> 0 *)
   ok :: "[i, i] => o"     (infixl \<open>ok\<close> 65)  where
-    "F ok G \<equiv> Acts(F) \<subseteq> AllowedActs(G) &
+    "F ok G \<equiv> Acts(F) \<subseteq> AllowedActs(G) \<and>
                Acts(G) \<subseteq> AllowedActs(F)"
 
 definition
@@ -37,8 +37,8 @@
 definition
   (*Characterizes safety properties.  Used with specifying AllowedActs*)
   safety_prop :: "i => o"  where
-  "safety_prop(X) \<equiv> X\<subseteq>program &
-      SKIP \<in> X & (\<forall>G \<in> program. Acts(G) \<subseteq> (\<Union>F \<in> X. Acts(F)) \<longrightarrow> G \<in> X)"
+  "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)
@@ -71,7 +71,7 @@
 
 subsection\<open>SKIP and safety properties\<close>
 
-lemma SKIP_in_constrains_iff [iff]: "(SKIP \<in> A co B) \<longleftrightarrow> (A\<subseteq>B & st_set(A))"
+lemma SKIP_in_constrains_iff [iff]: "(SKIP \<in> A co B) \<longleftrightarrow> (A\<subseteq>B \<and> st_set(A))"
 by (unfold constrains_def st_set_def, auto)
 
 lemma SKIP_in_Constrains_iff [iff]: "(SKIP \<in> A Co B)\<longleftrightarrow> (state \<inter> A\<subseteq>B)"
@@ -231,12 +231,12 @@
 done
 
 lemma Join_constrains [iff]:
-     "(F \<squnion> G \<in> A co B) \<longleftrightarrow> (programify(F) \<in> A co B & programify(G) \<in> A co B)"
+     "(F \<squnion> G \<in> A co B) \<longleftrightarrow> (programify(F) \<in> A co B \<and> programify(G) \<in> A co B)"
 by (auto simp add: constrains_def)
 
 lemma Join_unless [iff]:
      "(F \<squnion> G \<in> A unless B) \<longleftrightarrow>
-    (programify(F) \<in> A unless B & programify(G) \<in> A unless B)"
+    (programify(F) \<in> A unless B \<and> programify(G) \<in> A unless B)"
 by (simp add: Join_constrains unless_def)
 
 
@@ -247,7 +247,7 @@
 lemma Join_constrains_weaken:
      "\<lbrakk>F \<in> A co A';  G \<in> B co B'\<rbrakk>
       \<Longrightarrow> F \<squnion> G \<in> (A \<inter> B) co (A' \<union> B')"
-apply (subgoal_tac "st_set (A) & st_set (B) & F \<in> program & G \<in> program")
+apply (subgoal_tac "st_set (A) \<and> st_set (B) \<and> F \<in> program \<and> G \<in> program")
 prefer 2 apply (blast dest: constrainsD2, simp)
 apply (blast intro: constrains_weaken)
 done
@@ -267,7 +267,7 @@
 done
 
 lemma JOIN_stable:
-     "(\<Squnion>i \<in> I. F(i)) \<in>  stable(A) \<longleftrightarrow> ((\<forall>i \<in> I. programify(F(i)) \<in> stable(A)) & st_set(A))"
+     "(\<Squnion>i \<in> I. F(i)) \<in>  stable(A) \<longleftrightarrow> ((\<forall>i \<in> I. programify(F(i)) \<in> stable(A)) \<and> st_set(A))"
 apply (auto simp add: stable_def constrains_def JOIN_def)
 apply (cut_tac F = "F (i) " in Acts_type)
 apply (drule_tac x = act in bspec, auto)
@@ -298,7 +298,7 @@
 
 lemma Join_stable [iff]:
      " (F \<squnion> G \<in> stable(A)) \<longleftrightarrow>
-      (programify(F) \<in> stable(A) & programify(G) \<in>  stable(A))"
+      (programify(F) \<in> stable(A) \<and> programify(G) \<in>  stable(A))"
 by (simp add: stable_def)
 
 lemma initially_JoinI [intro!]:
@@ -308,7 +308,7 @@
 lemma invariant_JoinI:
      "\<lbrakk>F \<in> invariant(A); G \<in> invariant(A)\<rbrakk>
       \<Longrightarrow> F \<squnion> G \<in> invariant(A)"
-apply (subgoal_tac "F \<in> program&G \<in> program")
+apply (subgoal_tac "F \<in> program\<and>G \<in> program")
 prefer 2 apply (blast dest: invariantD2)
 apply (simp add: invariant_def)
 apply (auto intro: Join_in_program)
@@ -347,14 +347,14 @@
 lemma JOIN_ensures:
      "i \<in> I \<Longrightarrow>
       (\<Squnion>i \<in> I. F(i)) \<in> A ensures B \<longleftrightarrow>
-      ((\<forall>i \<in> I. programify(F(i)) \<in> (A-B) co (A \<union> B)) &
+      ((\<forall>i \<in> I. programify(F(i)) \<in> (A-B) co (A \<union> B)) \<and>
       (\<exists>i \<in> I. programify(F(i)) \<in> A ensures B))"
 by (auto simp add: ensures_def JOIN_constrains JOIN_transient)
 
 
 lemma Join_ensures:
      "F \<squnion> G \<in> A ensures B  \<longleftrightarrow>
-      (programify(F) \<in> (A-B) co (A \<union> B) & programify(G) \<in> (A-B) co (A \<union> B) &
+      (programify(F) \<in> (A-B) co (A \<union> B) \<and> programify(G) \<in> (A-B) co (A \<union> B) \<and>
        (programify(F) \<in>  transient (A-B) | programify(G) \<in> transient (A-B)))"
 
 apply (unfold ensures_def)
@@ -373,7 +373,7 @@
    weaker than G \<in> stable A *)
 lemma stable_Join_Always1:
      "\<lbrakk>F \<in> stable(A);  G \<in> invariant(A)\<rbrakk> \<Longrightarrow> F \<squnion> G \<in> Always(A)"
-apply (subgoal_tac "F \<in> program & G \<in> program & st_set (A) ")
+apply (subgoal_tac "F \<in> program \<and> G \<in> program \<and> st_set (A) ")
 prefer 2 apply (blast dest: invariantD2 stableD2)
 apply (simp add: Always_def invariant_def initially_def Stable_eq_stable)
 apply (force intro: stable_Int)
@@ -390,7 +390,7 @@
 
 lemma stable_Join_ensures1:
      "\<lbrakk>F \<in> stable(A);  G \<in> A ensures B\<rbrakk> \<Longrightarrow> F \<squnion> G \<in> A ensures B"
-apply (subgoal_tac "F \<in> program & G \<in> program & st_set (A) ")
+apply (subgoal_tac "F \<in> program \<and> G \<in> program \<and> st_set (A) ")
 prefer 2 apply (blast dest: stableD2 ensures_type [THEN subsetD])
 apply (simp (no_asm_simp) add: Join_ensures)
 apply (simp add: stable_def ensures_def)
@@ -414,7 +414,7 @@
 by (auto dest: Acts_type [THEN subsetD] simp add: ok_def)
 
 lemma ok_Join_commute:
-     "(F ok G & (F \<squnion> G) ok H) \<longleftrightarrow> (G ok H & F ok (G \<squnion> H))"
+     "(F ok G \<and> (F \<squnion> G) ok H) \<longleftrightarrow> (G ok H \<and> F ok (G \<squnion> H))"
 by (auto simp add: ok_def)
 
 lemma ok_commute: "(F ok G) \<longleftrightarrow>(G ok F)"
@@ -422,14 +422,14 @@
 
 lemmas ok_sym = ok_commute [THEN iffD1]
 
-lemma ok_iff_OK: "OK({<0,F>,<1,G>,<2,H>}, snd) \<longleftrightarrow> (F ok G & (F \<squnion> G) ok H)"
+lemma ok_iff_OK: "OK({<0,F>,<1,G>,<2,H>}, 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+)
 
-lemma ok_Join_iff1 [iff]: "F ok (G \<squnion> H) \<longleftrightarrow> (F ok G & F ok H)"
+lemma ok_Join_iff1 [iff]: "F ok (G \<squnion> H) \<longleftrightarrow> (F ok G \<and> F ok H)"
 by (auto simp add: ok_def)
 
-lemma ok_Join_iff2 [iff]: "(G \<squnion> H) ok F \<longleftrightarrow> (G ok F & H ok F)"
+lemma ok_Join_iff2 [iff]: "(G \<squnion> H) ok F \<longleftrightarrow> (G ok F \<and> H ok F)"
 by (auto simp add: ok_def)
 
 (*useful?  Not with the previous two around*)
@@ -456,7 +456,7 @@
 
 lemma OK_cons_iff:
      "OK(cons(i, I), F) \<longleftrightarrow>
-      (i \<in> I & OK(I, F)) | (i\<notin>I & OK(I, F) & F(i) ok JOIN(I,F))"
+      (i \<in> I \<and> OK(I, F)) | (i\<notin>I \<and> OK(I, F) \<and> F(i) ok JOIN(I,F))"
 apply (simp add: OK_iff_ok)
 apply (blast intro: ok_sym)
 done
@@ -480,7 +480,7 @@
 done
 
 lemma ok_iff_Allowed:
-     "F ok G \<longleftrightarrow> (programify(F) \<in> Allowed(programify(G)) &
+     "F ok G \<longleftrightarrow> (programify(F) \<in> Allowed(programify(G)) \<and>
    programify(G) \<in> Allowed(programify(F)))"
 by (simp add: ok_def Allowed_def)
 
@@ -525,7 +525,7 @@
 
 (*For safety_prop to hold, the property must be satisfiable!*)
 lemma safety_prop_constrains [iff]:
-     "safety_prop(A co B) \<longleftrightarrow> (A \<subseteq> B & st_set(A))"
+     "safety_prop(A co B) \<longleftrightarrow> (A \<subseteq> B \<and> st_set(A))"
 by (simp add: safety_prop_def constrains_def st_set_def, blast)
 
 (* To be used with resolution *)
@@ -566,7 +566,7 @@
 
 lemma def_UNION_ok_iff:
 "\<lbrakk>F \<equiv> mk_program(init,acts, \<Union>G \<in> X. Acts(G)); safety_prop(X)\<rbrakk>
-      \<Longrightarrow> F ok G \<longleftrightarrow> (programify(G) \<in> X & acts \<inter> Pow(state*state) \<subseteq> AllowedActs(G))"
+      \<Longrightarrow> F ok G \<longleftrightarrow> (programify(G) \<in> X \<and> acts \<inter> Pow(state*state) \<subseteq> AllowedActs(G))"
 apply (unfold ok_def)
 apply (drule_tac G = G in safety_prop_Acts_iff)
 apply (cut_tac F = G in AllowedActs_type)
--- a/src/ZF/UNITY/WFair.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/UNITY/WFair.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -17,8 +17,8 @@
   (* This definition specifies weak fairness.  The rest of the theory
     is generic to all forms of fairness.*)
   transient :: "i=>i"  where
-  "transient(A) \<equiv>{F \<in> program. (\<exists>act\<in>Acts(F). A<=domain(act) &
-                               act``A \<subseteq> state-A) & st_set(A)}"
+  "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
@@ -65,7 +65,7 @@
 by (unfold transient_def, auto)
 
 lemma transientD2:
-"F \<in> transient(A) \<Longrightarrow> F \<in> program & st_set(A)"
+"F \<in> transient(A) \<Longrightarrow> F \<in> program \<and> st_set(A)"
 apply (unfold transient_def, auto)
 done
 
@@ -127,7 +127,7 @@
 apply (auto simp add: ensures_def transient_type [THEN subsetD])
 done
 
-lemma ensuresD: "F \<in> A ensures B \<Longrightarrow> F:(A-B) co (A \<union> B) & F \<in> transient (A-B)"
+lemma ensuresD: "F \<in> A ensures B \<Longrightarrow> F:(A-B) co (A \<union> B) \<and> F \<in> transient (A-B)"
 by (unfold ensures_def, auto)
 
 lemma ensures_weaken_R: "\<lbrakk>F \<in> A ensures A'; A'<=B'\<rbrakk> \<Longrightarrow> F \<in> A ensures B'"
@@ -164,7 +164,7 @@
 by (unfold leadsTo_def, auto)
 
 lemma leadsToD2:
-"F \<in> A \<longmapsto> B \<Longrightarrow> F \<in> program & st_set(A) & st_set(B)"
+"F \<in> A \<longmapsto> B \<Longrightarrow> F \<in> program \<and> st_set(A) \<and> st_set(B)"
 apply (unfold leadsTo_def st_set_def)
 apply (blast dest: leads_left leads_right)
 done
@@ -240,14 +240,14 @@
 lemma leadsTo_refl: "\<lbrakk>F \<in> program; st_set(A)\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto> A"
 by (blast intro: subset_imp_leadsTo)
 
-lemma leadsTo_refl_iff: "F \<in> A \<longmapsto> A \<longleftrightarrow> F \<in> program & st_set(A)"
+lemma leadsTo_refl_iff: "F \<in> A \<longmapsto> A \<longleftrightarrow> F \<in> program \<and> st_set(A)"
 by (auto intro: leadsTo_refl dest: leadsToD2)
 
-lemma empty_leadsTo: "F \<in> 0 \<longmapsto> B \<longleftrightarrow> (F \<in> program & st_set(B))"
+lemma empty_leadsTo: "F \<in> 0 \<longmapsto> B \<longleftrightarrow> (F \<in> program \<and> st_set(B))"
 by (auto intro: subset_imp_leadsTo dest: leadsToD2)
 declare empty_leadsTo [iff]
 
-lemma leadsTo_state: "F \<in> A \<longmapsto> state \<longleftrightarrow> (F \<in> program & st_set(A))"
+lemma leadsTo_state: "F \<in> A \<longmapsto> state \<longleftrightarrow> (F \<in> program \<and> st_set(A))"
 by (auto intro: subset_imp_leadsTo dest: leadsToD2 st_setD)
 declare leadsTo_state [iff]
 
@@ -272,15 +272,15 @@
 done
 
 (*Distributes over binary unions*)
-lemma leadsTo_Un_distrib: "F:(A \<union> B) \<longmapsto> C  \<longleftrightarrow>  (F \<in> A \<longmapsto> C & F \<in> B \<longmapsto> C)"
+lemma leadsTo_Un_distrib: "F:(A \<union> B) \<longmapsto> C  \<longleftrightarrow>  (F \<in> A \<longmapsto> C \<and> F \<in> B \<longmapsto> C)"
 by (blast intro: leadsTo_Un leadsTo_weaken_L)
 
 lemma leadsTo_UN_distrib:
-"(F:(\<Union>i \<in> I. A(i)) \<longmapsto> B)\<longleftrightarrow> ((\<forall>i \<in> I. F \<in> A(i) \<longmapsto> B) & F \<in> program & st_set(B))"
+"(F:(\<Union>i \<in> I. A(i)) \<longmapsto> B)\<longleftrightarrow> ((\<forall>i \<in> I. F \<in> A(i) \<longmapsto> B) \<and> F \<in> program \<and> st_set(B))"
 apply (blast dest: leadsToD2 intro: leadsTo_UN leadsTo_weaken_L)
 done
 
-lemma leadsTo_Union_distrib: "(F \<in> \<Union>(S) \<longmapsto> B) \<longleftrightarrow>  (\<forall>A \<in> S. F \<in> A \<longmapsto> B) & F \<in> program & st_set(B)"
+lemma leadsTo_Union_distrib: "(F \<in> \<Union>(S) \<longmapsto> B) \<longleftrightarrow>  (\<forall>A \<in> S. F \<in> A \<longmapsto> B) \<and> F \<in> program \<and> st_set(B)"
 by (blast dest: leadsToD2 intro: leadsTo_Union leadsTo_weaken_L)
 
 text\<open>Set difference: maybe combine with \<open>leadsTo_weaken_L\<close>??\<close>
@@ -298,14 +298,14 @@
 
 (*Binary union version*)
 lemma leadsTo_Un_Un: "\<lbrakk>F \<in> A \<longmapsto> A'; F \<in> B \<longmapsto> B'\<rbrakk> \<Longrightarrow> F \<in> (A \<union> B) \<longmapsto> (A' \<union> B')"
-apply (subgoal_tac "st_set (A) & st_set (A') & st_set (B) & st_set (B') ")
+apply (subgoal_tac "st_set (A) \<and> st_set (A') \<and> st_set (B) \<and> st_set (B') ")
 prefer 2 apply (blast dest: leadsToD2)
 apply (blast intro: leadsTo_Un leadsTo_weaken_R)
 done
 
 (** The cancellation law **)
 lemma leadsTo_cancel2: "\<lbrakk>F \<in> A \<longmapsto> (A' \<union> B); F \<in> B \<longmapsto> B'\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto> (A' \<union> B')"
-apply (subgoal_tac "st_set (A) & st_set (A') & st_set (B) & st_set (B') &F \<in> program")
+apply (subgoal_tac "st_set (A) \<and> st_set (A') \<and> st_set (B) \<and> st_set (B') \<and>F \<in> program")
 prefer 2 apply (blast dest: leadsToD2)
 apply (blast intro: leadsTo_Trans leadsTo_Un_Un leadsTo_refl)
 done
@@ -397,9 +397,9 @@
   "\<lbrakk>F \<in> za \<longmapsto> zb;
       P(zb);
       \<And>A B. \<lbrakk>F \<in> A ensures B;  F \<in> B \<longmapsto> zb;  P(B); st_set(A)\<rbrakk> \<Longrightarrow> P(A);
-      \<And>S. \<forall>A \<in> S. F \<in> A \<longmapsto> zb & P(A) & st_set(A) \<Longrightarrow> P(\<Union>(S))
+      \<And>S. \<forall>A \<in> S. F \<in> A \<longmapsto> zb \<and> P(A) \<and> st_set(A) \<Longrightarrow> P(\<Union>(S))
 \<rbrakk> \<Longrightarrow> P(za)"
-apply (subgoal_tac " (F \<in> za \<longmapsto> zb) & P (za) ")
+apply (subgoal_tac " (F \<in> za \<longmapsto> zb) \<and> P (za) ")
 apply (erule conjunct2)
 apply (frule leadsToD2)
 apply (erule leadsTo_induct_pre_aux)
@@ -449,7 +449,7 @@
 
 lemma psp:
 "\<lbrakk>F \<in> A \<longmapsto> A'; F \<in> B co B'; st_set(B')\<rbrakk>\<Longrightarrow> F:(A \<inter> B') \<longmapsto> ((A' \<inter> B) \<union> (B' - B))"
-apply (subgoal_tac "F \<in> program & st_set (A) & st_set (A') & st_set (B) ")
+apply (subgoal_tac "F \<in> program \<and> st_set (A) \<and> st_set (A') \<and> st_set (B) ")
 prefer 2 apply (blast dest!: constrainsD2 leadsToD2)
 apply (erule leadsTo_induct)
 prefer 3 apply (blast intro: leadsTo_Union_Int)
@@ -471,7 +471,7 @@
    "\<lbrakk>F \<in> A \<longmapsto> A';  F \<in> B unless B'; st_set(B); st_set(B')\<rbrakk>
     \<Longrightarrow> F \<in> (A \<inter> B) \<longmapsto> ((A' \<inter> B) \<union> B')"
 apply (unfold unless_def)
-apply (subgoal_tac "st_set (A) &st_set (A') ")
+apply (subgoal_tac "st_set (A) \<and>st_set (A') ")
 prefer 2 apply (blast dest: leadsToD2)
 apply (drule psp, assumption, blast)
 apply (blast intro: leadsTo_weaken)
@@ -556,7 +556,7 @@
 done
 declare wlt_st_set [iff]
 
-lemma wlt_leadsTo_iff: "F \<in> wlt(F, B) \<longmapsto> B \<longleftrightarrow> (F \<in> program & st_set(B))"
+lemma wlt_leadsTo_iff: "F \<in> wlt(F, B) \<longmapsto> B \<longleftrightarrow> (F \<in> program \<and> st_set(B))"
 apply (unfold wlt_def)
 apply (blast dest: leadsToD2 intro!: leadsTo_Union)
 done
@@ -571,7 +571,7 @@
 done
 
 (*Misra's property W2*)
-lemma leadsTo_eq_subset_wlt: "F \<in> A \<longmapsto> B \<longleftrightarrow> (A \<subseteq> wlt(F,B) & F \<in> program & st_set(B))"
+lemma leadsTo_eq_subset_wlt: "F \<in> A \<longmapsto> B \<longleftrightarrow> (A \<subseteq> wlt(F,B) \<and> F \<in> program \<and> st_set(B))"
 apply auto
 apply (blast dest: leadsToD2 leadsTo_subset intro: leadsTo_weaken_L wlt_leadsTo)+
 done
@@ -594,7 +594,7 @@
 (*Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"*)
 (* slightly different from the HOL one \<in> B here is bounded *)
 lemma leadsTo_123: "F \<in> A \<longmapsto> A'
-      \<Longrightarrow> \<exists>B \<in> Pow(state). A<=B & F \<in> B \<longmapsto> A' & F \<in> (B-A') co (B \<union> A')"
+      \<Longrightarrow> \<exists>B \<in> Pow(state). A<=B \<and> F \<in> B \<longmapsto> A' \<and> F \<in> (B-A') co (B \<union> A')"
 apply (frule leadsToD2)
 apply (erule leadsTo_induct)
   txt\<open>Basis\<close>
@@ -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 & F \<in> Ba \<longmapsto> B & F \<in> Ba - B co Ba \<union> B}) ")
+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}) ")
 defer 1
 apply (rule AC_ball_Pi, safe)
 apply (rotate_tac 1)
@@ -634,7 +634,7 @@
        F \<in> A \<longmapsto> (A' \<union> C);  F \<in> A' co (A' \<union> C);
        F \<in> B \<longmapsto> (B' \<union> C);  F \<in> B' co (B' \<union> C)\<rbrakk>
     \<Longrightarrow> F \<in> (A \<inter> B) \<longmapsto> ((A' \<inter> B') \<union> C)"
-apply (subgoal_tac "st_set (C) &st_set (W) &st_set (W-C) &st_set (A') &st_set (A) & st_set (B) & st_set (B') & F \<in> program")
+apply (subgoal_tac "st_set (C) \<and>st_set (W) \<and>st_set (W-C) \<and>st_set (A') \<and>st_set (A) \<and> st_set (B) \<and> st_set (B') \<and> F \<in> program")
  prefer 2
  apply simp
  apply (blast dest!: leadsToD2)
--- a/src/ZF/Univ.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Univ.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -283,7 +283,7 @@
 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 & b: C"
+lemma Transset_Pair_subset: "\<lbrakk><a,b> \<subseteq> C; Transset(C)\<rbrakk> \<Longrightarrow> a: C \<and> b: C"
 by (unfold Pair_def Transset_def, blast)
 
 lemma Transset_Pair_subset_VLimit:
@@ -317,7 +317,7 @@
 lemma in_VLimit:
   "\<lbrakk>a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i);  Limit(i);
       \<And>x y j. \<lbrakk>j<i; 1:j; x \<in> Vfrom(A,j); y \<in> Vfrom(A,j)\<rbrakk>
-               \<Longrightarrow> \<exists>k. h(x,y) \<in> Vfrom(A,k) & k<i\<rbrakk>
+               \<Longrightarrow> \<exists>k. h(x,y) \<in> Vfrom(A,k) \<and> k<i\<rbrakk>
    \<Longrightarrow> h(a,b) \<in> Vfrom(A,i)"
 txt\<open>Infer that a, b occur at ordinals x,xa < i.\<close>
 apply (erule Limit_VfromE, assumption)
@@ -677,7 +677,7 @@
 subsubsection\<open>Closure under Finite Powerset\<close>
 
 lemma Fin_Vfrom_lemma:
-     "\<lbrakk>b: Fin(Vfrom(A,i));  Limit(i)\<rbrakk> \<Longrightarrow> \<exists>j. b \<subseteq> Vfrom(A,j) & j<i"
+     "\<lbrakk>b: Fin(Vfrom(A,i));  Limit(i)\<rbrakk> \<Longrightarrow> \<exists>j. b \<subseteq> Vfrom(A,j) \<and> j<i"
 apply (erule Fin_induct)
 apply (blast dest!: Limit_has_0, safe)
 apply (erule Limit_VfromE, assumption)
@@ -759,7 +759,7 @@
 text\<open>This version says a, b exist one level down, in the smaller set Vfrom(X,i)\<close>
 lemma doubleton_in_Vfrom_D:
      "\<lbrakk>{a,b} \<in> Vfrom(X,succ(i));  Transset(X)\<rbrakk>
-      \<Longrightarrow> a \<in> Vfrom(X,i)  &  b \<in> Vfrom(X,i)"
+      \<Longrightarrow> a \<in> Vfrom(X,i)  \<and>  b \<in> Vfrom(X,i)"
 by (drule Transset_Vfrom_succ [THEN equalityD1, THEN subsetD, THEN PowD],
     assumption, fast)
 
@@ -776,7 +776,7 @@
 
 lemma Pair_in_Vfrom_D:
     "\<lbrakk><a,b> \<in> Vfrom(X,succ(i));  Transset(X)\<rbrakk>
-     \<Longrightarrow> a \<in> Vfrom(X,i)  &  b \<in> Vfrom(X,i)"
+     \<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)
 done
--- a/src/ZF/ZF_Base.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/ZF_Base.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -49,7 +49,7 @@
    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)) & P(x,y))"
+  where "Replace(A,P) \<equiv> PrimReplace(A, %x y. (\<exists>!z. P(x,z)) \<and> P(x,y))"
 
 syntax
   "_Replace"  :: "[pttrn, pttrn, i, o] => i"  (\<open>(1{_ ./ _ \<in> _, _})\<close>)
@@ -71,7 +71,7 @@
 (* Separation and Pairing can be derived from the Replacement
    and Powerset Axioms using the following definitions. *)
 definition Collect :: "[i, i \<Rightarrow> o] \<Rightarrow> i"
-  where "Collect(A,P) \<equiv> {y . x\<in>A, x=y & P(x)}"
+  where "Collect(A,P) \<equiv> {y . x\<in>A, x=y \<and> P(x)}"
 
 syntax
   "_Collect" :: "[pttrn, i, o] \<Rightarrow> i"  (\<open>(1{_ \<in> _ ./ _})\<close>)
@@ -98,7 +98,7 @@
   set enumerations translate as {a,...,z} = cons(a,...,cons(z,0)...)*)
 
 definition Upair :: "[i, i] => i"
-  where "Upair(a,b) \<equiv> {y. x\<in>Pow(Pow(0)), (x=0 & y=a) | (x=Pow(0) & y=b)}"
+  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>
   where subset_def: "A \<subseteq> B \<equiv> \<forall>x\<in>A. x\<in>B"
@@ -157,7 +157,7 @@
   where the_def: "The(P)    \<equiv> \<Union>({y . x \<in> {0}, P(y)})"
 
 definition If :: "[o, i, i] \<Rightarrow> i"  (\<open>(if (_)/ then (_)/ else (_))\<close> [10] 10)
-  where if_def: "if P then a else b \<equiv> THE z. P & z=a | \<not>P & z=b"
+  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>)
@@ -240,7 +240,7 @@
   where "f`a \<equiv> \<Union>(f``{a})"
 
 definition Pi :: "[i, i \<Rightarrow> i] \<Rightarrow> i"
-  where "Pi(A,B) \<equiv> {f\<in>Pow(Sigma(A,B)). A\<subseteq>domain(f) & function(f)}"
+  where "Pi(A,B) \<equiv> {f\<in>Pow(Sigma(A,B)). A\<subseteq>domain(f) \<and> function(f)}"
 
 abbreviation function_space :: "[i, i] \<Rightarrow> i"  (infixr \<open>\<rightarrow>\<close> 60)  \<comment> \<open>function space\<close>
   where "A \<rightarrow> B \<equiv> Pi(A, \<lambda>_. B)"
@@ -347,7 +347,7 @@
 by (simp add: Bex_def, blast)
 
 (*We do not even have @{term"(\<exists>x\<in>A. True) <-> True"} unless @{term"A" is nonempty\<And>*)
-lemma bex_triv [simp]: "(\<exists>x\<in>A. P) <-> ((\<exists>x. x\<in>A) & P)"
+lemma bex_triv [simp]: "(\<exists>x\<in>A. P) <-> ((\<exists>x. x\<in>A) \<and> P)"
 by (simp add: Bex_def)
 
 lemma bex_cong [cong]:
@@ -429,7 +429,7 @@
 subsection\<open>Rules for Replace -- the derived form of replacement\<close>
 
 lemma Replace_iff:
-    "b \<in> {y. x\<in>A, P(x,y)}  <->  (\<exists>x\<in>A. P(x,b) & (\<forall>y. P(x,y) \<longrightarrow> y=b))"
+    "b \<in> {y. x\<in>A, P(x,y)}  <->  (\<exists>x\<in>A. P(x,b) \<and> (\<forall>y. P(x,y) \<longrightarrow> y=b))"
 apply (unfold Replace_def)
 apply (rule replacement [THEN iff_trans], blast+)
 done
@@ -493,7 +493,7 @@
 subsection\<open>Rules for Collect -- forming a subset by separation\<close>
 
 (*Separation is derivable from Replacement*)
-lemma separation [simp]: "a \<in> {x\<in>A. P(x)} <-> a\<in>A & P(a)"
+lemma separation [simp]: "a \<in> {x\<in>A. P(x)} <-> a\<in>A \<and> P(a)"
 by (unfold Collect_def, blast)
 
 lemma CollectI [intro!]: "\<lbrakk>a\<in>A;  P(a)\<rbrakk> \<Longrightarrow> a \<in> {x\<in>A. P(x)}"
@@ -586,7 +586,7 @@
 subsection\<open>Rules for Inter\<close>
 
 (*Not obviously useful for proving InterI, InterD, InterE*)
-lemma Inter_iff: "A \<in> \<Inter>(C) <-> (\<forall>x\<in>C. A: x) & C\<noteq>0"
+lemma Inter_iff: "A \<in> \<Inter>(C) <-> (\<forall>x\<in>C. A: x) \<and> C\<noteq>0"
 by (simp add: Inter_def Ball_def, blast)
 
 (* Intersection is well-behaved only if the family is non-empty! *)
@@ -609,7 +609,7 @@
 
 (* @{term"\<Inter>x\<in>A. B(x)"} abbreviates @{term"\<Inter>({B(x). x\<in>A})"} *)
 
-lemma INT_iff: "b \<in> (\<Inter>x\<in>A. B(x)) <-> (\<forall>x\<in>A. b \<in> B(x)) & A\<noteq>0"
+lemma INT_iff: "b \<in> (\<Inter>x\<in>A. B(x)) <-> (\<forall>x\<in>A. b \<in> B(x)) \<and> A\<noteq>0"
 by (force simp add: Inter_def)
 
 lemma INT_I: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b: B(x);  A\<noteq>0\<rbrakk> \<Longrightarrow> b: (\<Inter>x\<in>A. B(x))"
--- a/src/ZF/Zorn.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/Zorn.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -12,7 +12,7 @@
 
 definition
   Subset_rel :: "i=>i"  where
-   "Subset_rel(A) \<equiv> {z \<in> A*A . \<exists>x y. z=<x,y> & x<=y & x\<noteq>y}"
+   "Subset_rel(A) \<equiv> {z \<in> A*A . \<exists>x y. z=<x,y> \<and> x<=y \<and> x\<noteq>y}"
 
 definition
   chain      :: "i=>i"  where
@@ -20,7 +20,7 @@
 
 definition
   super      :: "[i,i]=>i"  where
-   "super(A,c)    \<equiv> {d \<in> chain(A). c<=d & c\<noteq>d}"
+   "super(A,c)    \<equiv> {d \<in> chain(A). c<=d \<and> c\<noteq>d}"
 
 definition
   maxchain   :: "i=>i"  where
--- a/src/ZF/equalities.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/equalities.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -22,7 +22,7 @@
   The following are not added to the default simpset because
   (a) they duplicate the body and (b) there are no similar rules for \<open>Int\<close>.\<close>
 
-lemma ball_Un: "(\<forall>x \<in> A\<union>B. P(x)) \<longleftrightarrow> (\<forall>x \<in> A. P(x)) & (\<forall>x \<in> B. P(x))"
+lemma ball_Un: "(\<forall>x \<in> A\<union>B. P(x)) \<longleftrightarrow> (\<forall>x \<in> A. P(x)) \<and> (\<forall>x \<in> B. P(x))"
   by blast
 
 lemma bex_Un: "(\<exists>x \<in> A\<union>B. P(x)) \<longleftrightarrow> (\<exists>x \<in> A. P(x)) | (\<exists>x \<in> B. P(x))"
@@ -76,7 +76,7 @@
 lemma subset_consI: "B \<subseteq> cons(a,B)"
 by blast
 
-lemma cons_subset_iff [iff]: "cons(a,B)\<subseteq>C \<longleftrightarrow> a\<in>C & B\<subseteq>C"
+lemma cons_subset_iff [iff]: "cons(a,B)\<subseteq>C \<longleftrightarrow> a\<in>C \<and> B\<subseteq>C"
 by blast
 
 (*A safe special case of subset elimination, adding no new variables
@@ -86,7 +86,7 @@
 lemma subset_empty_iff: "A\<subseteq>0 \<longleftrightarrow> A=0"
 by blast
 
-lemma subset_cons_iff: "C\<subseteq>cons(a,B) \<longleftrightarrow> C\<subseteq>B | (a\<in>C & C-{a} \<subseteq> B)"
+lemma subset_cons_iff: "C\<subseteq>cons(a,B) \<longleftrightarrow> C\<subseteq>B | (a\<in>C \<and> C-{a} \<subseteq> B)"
 by blast
 
 (* cons_def refers to Upair; reversing the equality LOOPS in rewriting!*)
@@ -134,7 +134,7 @@
     "\<lbrakk>succ(i) \<subseteq> j;  \<lbrakk>i\<in>j;  i\<subseteq>j\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
 by (unfold succ_def, blast)
 
-lemma succ_subset_iff: "succ(a) \<subseteq> B \<longleftrightarrow> (a \<subseteq> B & a \<in> B)"
+lemma succ_subset_iff: "succ(a) \<subseteq> B \<longleftrightarrow> (a \<subseteq> B \<and> a \<in> B)"
 by (unfold succ_def, blast)
 
 
@@ -142,7 +142,7 @@
 
 (** Intersection is the greatest lower bound of two sets **)
 
-lemma Int_subset_iff: "C \<subseteq> A \<inter> B \<longleftrightarrow> C \<subseteq> A & C \<subseteq> B"
+lemma Int_subset_iff: "C \<subseteq> A \<inter> B \<longleftrightarrow> C \<subseteq> A \<and> C \<subseteq> B"
 by blast
 
 lemma Int_lower1: "A \<inter> B \<subseteq> A"
@@ -211,7 +211,7 @@
 
 (** Union is the least upper bound of two sets *)
 
-lemma Un_subset_iff: "A \<union> B \<subseteq> C \<longleftrightarrow> A \<subseteq> C & B \<subseteq> C"
+lemma Un_subset_iff: "A \<union> B \<subseteq> C \<longleftrightarrow> A \<subseteq> C \<and> B \<subseteq> C"
 by blast
 
 lemma Un_upper1: "A \<subseteq> A \<union> B"
@@ -259,7 +259,7 @@
 lemma subset_Un_iff2: "A\<subseteq>B \<longleftrightarrow> B \<union> A = B"
 by (blast elim!: equalityE)
 
-lemma Un_empty [iff]: "(A \<union> B = 0) \<longleftrightarrow> (A = 0 & B = 0)"
+lemma Un_empty [iff]: "(A \<union> B = 0) \<longleftrightarrow> (A = 0 \<and> B = 0)"
 by blast
 
 lemma Un_eq_Union: "A \<union> B = \<Union>({A, B})"
@@ -273,7 +273,7 @@
 lemma Diff_contains: "\<lbrakk>C\<subseteq>A;  C \<inter> B = 0\<rbrakk> \<Longrightarrow> C \<subseteq> A-B"
 by blast
 
-lemma subset_Diff_cons_iff: "B \<subseteq> A - cons(c,C)  \<longleftrightarrow>  B\<subseteq>A-C & c \<notin> B"
+lemma subset_Diff_cons_iff: "B \<subseteq> A - cons(c,C)  \<longleftrightarrow>  B\<subseteq>A-C \<and> c \<notin> B"
 by blast
 
 lemma Diff_cancel: "A - A = 0"
@@ -571,7 +571,7 @@
 by blast
 
 lemma times_subset_iff:
-     "(A'*B' \<subseteq> A*B) \<longleftrightarrow> (A' = 0 | B' = 0 | (A'\<subseteq>A) & (B'\<subseteq>B))"
+     "(A'*B' \<subseteq> A*B) \<longleftrightarrow> (A' = 0 | B' = 0 | (A'\<subseteq>A) \<and> (B'\<subseteq>B))"
 by blast
 
 lemma Int_Sigma_eq:
@@ -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 & P(<x,y>)})"
+     "{z \<in> Sigma(A,B). P(z)} `` C = (\<Union>x \<in> A. {y \<in> B(x). x \<in> C \<and> P(<x,y>)})"
 by blast
 
 lemma image_Int_subset: "r``(A \<inter> B) \<subseteq> (r``A) \<inter> (r``B)"
@@ -941,11 +941,11 @@
 by blast
 
 lemma Collect_Collect_eq [simp]:
-     "Collect(Collect(A,P), Q) = Collect(A, %x. P(x) & Q(x))"
+     "Collect(Collect(A,P), Q) = Collect(A, %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) & Q(x))"
+     "Collect(A,P) \<inter> Collect(A,Q) = Collect(A, %x. P(x) \<and> Q(x))"
 by blast
 
 lemma Collect_Union_eq [simp]:
@@ -961,7 +961,7 @@
 lemma Collect_disj_eq: "{x\<in>A. P(x) | Q(x)} = Collect(A, P) \<union> Collect(A, Q)"
 by blast
 
-lemma Collect_conj_eq: "{x\<in>A. P(x) & Q(x)} = Collect(A, P) \<inter> Collect(A, Q)"
+lemma Collect_conj_eq: "{x\<in>A. P(x) \<and> Q(x)} = Collect(A, P) \<inter> Collect(A, Q)"
 by blast
 
 lemmas subset_SIs = subset_refl cons_subsetI subset_consI
--- a/src/ZF/ex/CoUnit.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/ex/CoUnit.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -55,7 +55,7 @@
 
 inductive_cases Con2E: "Con2(x, y) \<in> counit2"
 
-lemma Con2_iff: "Con2(x, y) = Con2(x', y') \<longleftrightarrow> x = x' & y = y'"
+lemma Con2_iff: "Con2(x, y) = Con2(x', y') \<longleftrightarrow> x = x' \<and> y = y'"
   \<comment> \<open>Proving freeness results.\<close>
   by (fast elim!: counit2.free_elims)
 
--- a/src/ZF/ex/Commutation.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/ex/Commutation.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -1,5 +1,5 @@
 (*  Title:      ZF/ex/Commutation.thy
-    Author:     Tobias Nipkow & Sidi Ould Ehmety
+    Author:     Tobias Nipkow \<and> Sidi Ould Ehmety
     Copyright   1995  TU Muenchen
 
 Commutation theory for proving the Church Rosser theorem.
@@ -10,7 +10,7 @@
 definition
   square  :: "[i, i, i, i] => 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 & <c,x> \<in> u)))"
+    (\<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)))"
 
 definition
   commute :: "[i, i] => o" where
@@ -27,7 +27,7 @@
 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^* & <y,z> \<in> r^*))"
+                        (\<exists>z. <x,z> \<in> r^* \<and> <y,z> \<in> r^*))"
 
 definition
   confluent :: "i=>o" where
--- a/src/ZF/ex/Group.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/ex/Group.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -35,7 +35,7 @@
 
 definition
   m_inv :: "i => i => i" (\<open>inv\<index> _\<close> [81] 80) where
-  "inv\<^bsub>G\<^esub> x \<equiv> (THE y. y \<in> carrier(G) & y \<cdot>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub> & x \<cdot>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub>)"
+  "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)
   assumes m_closed [intro, simp]:
@@ -87,7 +87,7 @@
 
 locale group = monoid +
   assumes inv_ex:
-     "\<And>x. x \<in> carrier(G) \<Longrightarrow> \<exists>y \<in> carrier(G). y \<cdot> x = \<one> & x \<cdot> y = \<one>"
+     "\<And>x. x \<in> carrier(G) \<Longrightarrow> \<exists>y \<in> carrier(G). y \<cdot> x = \<one> \<and> x \<cdot> y = \<one>"
 
 lemma (in group) is_group [simp]: "group(G)" by (rule group_axioms)
 
@@ -133,7 +133,7 @@
     with x xG show "x \<cdot> \<one> = x" by simp
   qed
   have inv_ex:
-    "\<And>x. x \<in> carrier(G) \<Longrightarrow> \<exists>y \<in> carrier(G). y \<cdot> x = \<one> & x \<cdot> y = \<one>"
+    "\<And>x. x \<in> carrier(G) \<Longrightarrow> \<exists>y \<in> carrier(G). y \<cdot> x = \<one> \<and> x \<cdot> y = \<one>"
   proof -
     fix x
     assume x: "x \<in> carrier(G)"
@@ -143,7 +143,7 @@
       by (simp add: m_assoc [symmetric] l_inv r_one)
     with x y have r_inv: "x \<cdot> y = \<one>"
       by simp
-    from x y show "\<exists>y \<in> carrier(G). y \<cdot> x = \<one> & x \<cdot> y = \<one>"
+    from x y show "\<exists>y \<in> carrier(G). y \<cdot> x = \<one> \<and> x \<cdot> y = \<one>"
       by (fast intro: l_inv r_inv)
   qed
   show ?thesis
@@ -152,7 +152,7 @@
 qed
 
 lemma (in group) inv [simp]:
-  "x \<in> carrier(G) \<Longrightarrow> inv x \<in> carrier(G) & inv x \<cdot> x = \<one> & x \<cdot> inv x = \<one>"
+  "x \<in> carrier(G) \<Longrightarrow> inv x \<in> carrier(G) \<and> inv x \<cdot> x = \<one> \<and> x \<cdot> inv x = \<one>"
   apply (frule inv_ex)
   apply (unfold Bex_def m_inv_def)
   apply (erule exE)
@@ -681,7 +681,7 @@
 text\<open>Alternative characterization of normal subgroups\<close>
 lemma (in group) normal_inv_iff:
      "(N \<lhd> G) \<longleftrightarrow>
-      (subgroup(N,G) & (\<forall>x \<in> carrier(G). \<forall>h \<in> N. x \<cdot> h \<cdot> (inv x) \<in> N))"
+      (subgroup(N,G) \<and> (\<forall>x \<in> carrier(G). \<forall>h \<in> N. x \<cdot> h \<cdot> (inv x) \<in> N))"
       (is "_ \<longleftrightarrow> ?rhs")
 proof
   assume N: "N \<lhd> G"
@@ -739,7 +739,7 @@
   assume "\<exists>h\<in>H. y = x \<cdot> h"
     and x: "x \<in> carrier(G)"
     and sb: "subgroup(H,G)"
-  then obtain h' where h': "h' \<in> H & x \<cdot> h' = y" by blast
+  then obtain h' where h': "h' \<in> H \<and> x \<cdot> h' = y" by blast
   show "\<exists>h\<in>H. x = y \<cdot> h"
   proof
     show "x = y \<cdot> inv h'" using h' x sb
--- a/src/ZF/ex/LList.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/ex/LList.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -74,7 +74,7 @@
 inductive_cases LConsE: "LCons(a,l) \<in> llist(A)"
 
 (*Proving freeness results*)
-lemma LCons_iff: "LCons(a,l)=LCons(a',l') \<longleftrightarrow> a=a' & l=l'"
+lemma LCons_iff: "LCons(a,l)=LCons(a',l') \<longleftrightarrow> a=a' \<and> l=l'"
 by auto
 
 lemma LNil_LCons_iff: "\<not> LNil=LCons(a,l)"
--- a/src/ZF/ex/Limit.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/ex/Limit.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -30,23 +30,23 @@
 definition
   po  :: "i=>o"  where
     "po(D) \<equiv>
-     (\<forall>x \<in> set(D). rel(D,x,x)) &
+     (\<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).
-       rel(D,x,y) \<longrightarrow> rel(D,y,z) \<longrightarrow> rel(D,x,z)) &
+       rel(D,x,y) \<longrightarrow> rel(D,y,z) \<longrightarrow> rel(D,x,z)) \<and>
      (\<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
     (* Chains are object level functions nat->set(D) *)
-    "chain(D,X) \<equiv> X \<in> nat->set(D) & (\<forall>n \<in> nat. rel(D,X`n,X`(succ(n))))"
+    "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(D,X,x) \<equiv> x \<in> set(D) & (\<forall>n \<in> nat. rel(D,X`n,x))"
+    "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(D,X,x) \<equiv> isub(D,X,x) & (\<forall>y. isub(D,X,y) \<longrightarrow> rel(D,x,y))"
+    "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
@@ -54,15 +54,15 @@
 
 definition
   cpo :: "i=>o"  where
-    "cpo(D) \<equiv> po(D) & (\<forall>X. chain(D,X) \<longrightarrow> (\<exists>x. islub(D,X,x)))"
+    "cpo(D) \<equiv> po(D) \<and> (\<forall>X. chain(D,X) \<longrightarrow> (\<exists>x. islub(D,X,x)))"
 
 definition
   pcpo :: "i=>o"  where
-    "pcpo(D) \<equiv> cpo(D) & (\<exists>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y))"
+    "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(D) \<equiv> THE x. x \<in> set(D) & (\<forall>y \<in> set(D). rel(D,x,y))"
+    "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
@@ -97,16 +97,16 @@
 definition
   matrix :: "[i,i]=>o"  where
     "matrix(D,M) \<equiv>
-     M \<in> nat -> (nat -> set(D)) &
-     (\<forall>n \<in> nat. \<forall>m \<in> nat. rel(D,M`n`m,M`succ(n)`m)) &
-     (\<forall>n \<in> nat. \<forall>m \<in> nat. rel(D,M`n`m,M`n`succ(m))) &
+     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>
+     (\<forall>n \<in> nat. \<forall>m \<in> nat. rel(D,M`n`m,M`n`succ(m))) \<and>
      (\<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(D,E,e,p) \<equiv>
-     e \<in> cont(D,E) & p \<in> cont(E,D) &
-     p O e = id(set(D)) & rel(cf(E,E),e O p,id(set(E)))"
+     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
@@ -133,18 +133,18 @@
 definition
   subcpo    :: "[i,i]=>o"  where
     "subcpo(D,E) \<equiv>
-     set(D) \<subseteq> set(E) &
-     (\<forall>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y) \<longleftrightarrow> rel(E,x,y)) &
+     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(D,E) \<equiv> subcpo(D,E) & bot(E):set(D)"
+    "subpcpo(D,E) \<equiv> subcpo(D,E) \<and> bot(E):set(D)"
 
 definition
   emb_chain :: "[i,i]=>o"  where
     "emb_chain(DD,ee) \<equiv>
-     (\<forall>n \<in> nat. cpo(DD`n)) & (\<forall>n \<in> nat. emb(DD`n,DD`succ(n),ee`n))"
+     (\<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
@@ -181,12 +181,12 @@
 definition
   commute   :: "[i,i,i,i=>i]=>o"  where
     "commute(DD,ee,E,r) \<equiv>
-     (\<forall>n \<in> nat. emb(DD`n,E,r(n))) &
+     (\<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(E,G,r,f,t) \<equiv> emb(E,G,t) & (\<forall>n \<in> nat. f(n) = t O r(n))"
+    "mediating(E,G,r,f,t) \<equiv> emb(E,G,t) \<and> (\<forall>n \<in> nat. f(n) = t O r(n))"
 
 
 lemmas nat_linear_le = Ord_linear_le [OF nat_into_Ord nat_into_Ord]
@@ -342,7 +342,7 @@
 by (simp add: pcpo_def)
 
 lemma pcpo_bot_ex1:
-    "pcpo(D) \<Longrightarrow> \<exists>! x. x \<in> set(D) & (\<forall>y \<in> set(D). rel(D,x,y))"
+    "pcpo(D) \<Longrightarrow> \<exists>! x. x \<in> set(D) \<and> (\<forall>y \<in> set(D). rel(D,x,y))"
 apply (simp add: pcpo_def)
 apply (blast intro: cpo_antisym)
 done
@@ -901,7 +901,7 @@
 lemma projpair_p_cont: "projpair(D,E,e,p) \<Longrightarrow> p \<in> cont(E,D)"
 by (simp add: projpair_def)
 
-lemma projpair_ep_cont: "projpair(D,E,e,p) \<Longrightarrow> e \<in> cont(D,E) & p \<in> cont(E,D)"
+lemma projpair_ep_cont: "projpair(D,E,e,p) \<Longrightarrow> e \<in> cont(D,E) \<and> p \<in> cont(E,D)"
 by (simp add: projpair_def)
 
 lemma projpair_eq: "projpair(D,E,e,p) \<Longrightarrow> p O e = id(set(D))"
@@ -2293,7 +2293,7 @@
    mediating
     (Dinf(DD,ee),G,rho_emb(DD,ee),f,
      lub(cf(Dinf(DD,ee),G),
-         \<lambda>n \<in> nat. f(n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n)))) &
+         \<lambda>n \<in> nat. f(n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n)))) \<and>
    (\<forall>t. mediating(Dinf(DD,ee),G,rho_emb(DD,ee),f,t) \<longrightarrow>
      t = lub(cf(Dinf(DD,ee),G),
          \<lambda>n \<in> nat. f(n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n))))"
--- a/src/ZF/ex/NatSum.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/ex/NatSum.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -1,5 +1,5 @@
 (*  Title:      ZF/ex/NatSum.thy
-    Author:     Tobias Nipkow & Lawrence C Paulson
+    Author:     Tobias Nipkow \<and> Lawrence C Paulson
 
 A summation operator. sum(f,n+1) is the sum of all f(i), i=0...n.
 
--- a/src/ZF/ex/Primes.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/ex/Primes.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -9,12 +9,12 @@
 
 definition
   divides :: "[i,i]=>o"              (infixl \<open>dvd\<close> 50)  where
-    "m dvd n \<equiv> m \<in> nat & n \<in> nat & (\<exists>k \<in> nat. n = m#*k)"
+    "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(p,m,n) \<equiv> ((p dvd m) & (p dvd n))   &
-                       (\<forall>d\<in>nat. (d dvd m) & (d dvd n) \<longrightarrow> d dvd p)"
+    "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
@@ -28,12 +28,12 @@
   
 definition
   prime   :: i                \<comment> \<open>the set of prime numbers\<close>  where
-   "prime \<equiv> {p \<in> nat. 1<p & (\<forall>m \<in> nat. m dvd p \<longrightarrow> m=1 | m=p)}"
+   "prime \<equiv> {p \<in> nat. 1<p \<and> (\<forall>m \<in> nat. m dvd p \<longrightarrow> m=1 | m=p)}"
 
 
 subsection\<open>The Divides Relation\<close>
 
-lemma dvdD: "m dvd n \<Longrightarrow> m \<in> nat & n \<in> nat & (\<exists>k \<in> nat. n = m#*k)"
+lemma dvdD: "m dvd n \<Longrightarrow> m \<in> nat \<and> n \<in> nat \<and> (\<exists>k \<in> nat. n = m#*k)"
 by (unfold divides_def, assumption)
 
 lemma dvdE:
@@ -176,7 +176,7 @@
 text\<open>Property 1: gcd(a,b) divides a and b\<close>
 
 lemma gcd_dvd_both:
-     "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> gcd (m, n) dvd m & gcd (m, n) dvd n"
+     "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> gcd (m, n) dvd m \<and> gcd (m, n) dvd n"
 apply (rule_tac m = m and n = n in gcd_induct)
 apply (simp_all add: gcd_non_0)
 apply (blast intro: dvd_mod_imp_dvd_raw nat_into_Ord [THEN Ord_0_lt])
@@ -217,7 +217,7 @@
 done
 
 lemma gcd_greatest_iff [simp]: "\<lbrakk>k \<in> nat; m \<in> nat; n \<in> nat\<rbrakk>  
-      \<Longrightarrow> (k dvd gcd (m, n)) \<longleftrightarrow> (k dvd m & k dvd n)"
+      \<Longrightarrow> (k dvd gcd (m, n)) \<longleftrightarrow> (k dvd m \<and> k dvd n)"
 by (blast intro!: gcd_greatest gcd_dvd1 gcd_dvd2 intro: dvd_trans)
 
 
@@ -387,7 +387,7 @@
 
 lemma reduction:
      "\<lbrakk>k#*k = p#*(j#*j); p \<in> prime; 0 < k; j \<in> nat; k \<in> nat\<rbrakk>  
-      \<Longrightarrow> k < p#*j & 0 < j"
+      \<Longrightarrow> k < p#*j \<and> 0 < j"
 apply (rule ccontr)
 apply (simp add: not_lt_iff_le prime_into_nat)
 apply (erule disjE)
--- a/src/ZF/ex/Ramsey.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/ex/Ramsey.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -36,17 +36,17 @@
 
 definition
   Clique  :: "[i,i,i]=>o" where
-    "Clique(C,V,E) \<equiv> (C \<subseteq> V) & (\<forall>x \<in> C. \<forall>y \<in> C. x\<noteq>y \<longrightarrow> <x,y> \<in> E)"
+    "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)"
 
 definition
   Indept  :: "[i,i,i]=>o" where
-    "Indept(I,V,E) \<equiv> (I \<subseteq> V) & (\<forall>x \<in> I. \<forall>y \<in> I. x\<noteq>y \<longrightarrow> <x,y> \<notin> E)"
+    "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)"
   
 definition
   Ramsey  :: "[i,i,i]=>o" where
-    "Ramsey(n,i,j) \<equiv> \<forall>V E. Symmetric(E) & Atleast(n,V) \<longrightarrow>  
-         (\<exists>C. Clique(C,V,E) & Atleast(i,C)) |       
-         (\<exists>I. Indept(I,V,E) & Atleast(j,I))"
+    "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))"
 
 (*** Cliques and Independent sets ***)
 
@@ -147,7 +147,7 @@
 lemma Indept_succ: 
     "\<lbrakk>Indept(I, {z \<in> V-{a}. <a,z> \<notin> E}, E);  Symmetric(E);  a \<in> V;   
         Atleast(j,I)\<rbrakk> \<Longrightarrow>    
-     Indept(cons(a,I), V, E) & Atleast(succ(j), cons(a,I))"
+     Indept(cons(a,I), V, E) \<and> Atleast(succ(j), cons(a,I))"
 apply (unfold Symmetric_def Indept_def)
 apply (blast intro!: Atleast_succI)
 done
@@ -156,7 +156,7 @@
 lemma Clique_succ: 
     "\<lbrakk>Clique(C, {z \<in> V-{a}. <a,z>:E}, E);  Symmetric(E);  a \<in> V;   
         Atleast(j,C)\<rbrakk> \<Longrightarrow>    
-     Clique(cons(a,C), V, E) & Atleast(succ(j), cons(a,C))"
+     Clique(cons(a,C), V, E) \<and> Atleast(succ(j), cons(a,C))"
 apply (unfold Symmetric_def Clique_def)
 apply (blast intro!: Atleast_succI)
 done
--- a/src/ZF/ex/Ring.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/ex/Ring.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -234,9 +234,9 @@
   ring_hom :: "[i,i] => i" where
   "ring_hom(R,S) \<equiv>
     {h \<in> carrier(R) -> carrier(S).
-      (\<forall>x y. x \<in> carrier(R) & y \<in> carrier(R) \<longrightarrow>
-        h ` (x \<cdot>\<^bsub>R\<^esub> y) = (h ` x) \<cdot>\<^bsub>S\<^esub> (h ` y) &
-        h ` (x \<oplus>\<^bsub>R\<^esub> y) = (h ` x) \<oplus>\<^bsub>S\<^esub> (h ` y)) &
+      (\<forall>x y. x \<in> carrier(R) \<and> y \<in> carrier(R) \<longrightarrow>
+        h ` (x \<cdot>\<^bsub>R\<^esub> y) = (h ` x) \<cdot>\<^bsub>S\<^esub> (h ` y) \<and>
+        h ` (x \<oplus>\<^bsub>R\<^esub> y) = (h ` x) \<oplus>\<^bsub>S\<^esub> (h ` y)) \<and>
       h ` \<one>\<^bsub>R\<^esub> = \<one>\<^bsub>S\<^esub>}"
 
 lemma ring_hom_memI:
--- a/src/ZF/ex/misc.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/ex/misc.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -36,15 +36,15 @@
 
 text\<open>These two are cited in Benzmueller and Kohlhase's system description of
  LEO, CADE-15, 1998 (page 139-143) as theorems LEO could not prove.\<close>
-lemma "(X = Y \<union> Z) \<longleftrightarrow> (Y \<subseteq> X & Z \<subseteq> X & (\<forall>V. Y \<subseteq> V & Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
+lemma "(X = Y \<union> Z) \<longleftrightarrow> (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
 by (blast intro!: equalityI)
 
 text\<open>the dual of the previous one\<close>
-lemma "(X = Y \<inter> Z) \<longleftrightarrow> (X \<subseteq> Y & X \<subseteq> Z & (\<forall>V. V \<subseteq> Y & V \<subseteq> Z \<longrightarrow> V \<subseteq> X))"
+lemma "(X = Y \<inter> Z) \<longleftrightarrow> (X \<subseteq> Y \<and> X \<subseteq> Z \<and> (\<forall>V. V \<subseteq> Y \<and> V \<subseteq> Z \<longrightarrow> V \<subseteq> X))"
 by (blast intro!: equalityI)
 
 text\<open>trivial example of term synthesis: apparently hard for some provers!\<close>
-schematic_goal "a \<noteq> b \<Longrightarrow> a:?X & b \<notin> ?X"
+schematic_goal "a \<noteq> b \<Longrightarrow> a:?X \<and> b \<notin> ?X"
 by blast
 
 text\<open>Nice blast benchmark.  Proved in 0.3s; old tactics can't manage it!\<close>
@@ -63,7 +63,7 @@
 by best
 
 text\<open>A characterization of functions suggested by Tobias Nipkow\<close>
-lemma "r \<in> domain(r)->B  \<longleftrightarrow>  r \<subseteq> domain(r)*B & (\<forall>X. r `` (r -`` X) \<subseteq> X)"
+lemma "r \<in> domain(r)->B  \<longleftrightarrow>  r \<subseteq> domain(r)*B \<and> (\<forall>X. r `` (r -`` X) \<subseteq> X)"
 by (unfold Pi_def function_def, best)
 
 
@@ -80,17 +80,17 @@
 (*Force helps prove conditions of rewrites such as comp_fun_apply, since
   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 & g \<in> B*B->B &  
+           {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>  
-       J \<in> hom(A,f,B,g) & K \<in> hom(B,g,C,h) \<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
 
 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 & g \<in> B*B->B &  
+           {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> J \<in> hom(A,f,B,g) & K \<in> hom(B,g,C,h) \<longrightarrow> (K O J) \<in> hom(A,f,C,h)"
+       \<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 16:51:35 2022 +0100
+++ b/src/ZF/func.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -20,12 +20,12 @@
 by (simp add: restrict_def relation_def, blast)
 
 lemma Pi_iff:
-    "f \<in> Pi(A,B) \<longleftrightarrow> function(f) & f<=Sigma(A,B) & A<=domain(f)"
+    "f \<in> Pi(A,B) \<longleftrightarrow> function(f) \<and> f<=Sigma(A,B) \<and> A<=domain(f)"
 by (unfold Pi_def, blast)
 
 (*For upward compatibility with the former definition*)
 lemma Pi_iff_old:
-    "f \<in> Pi(A,B) \<longleftrightarrow> f<=Sigma(A,B) & (\<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. <x,y>: f)"
 by (unfold Pi_def function_def, blast)
 
 lemma fun_is_function: "f \<in> Pi(A,B) \<Longrightarrow> function(f)"
@@ -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 & f`a = b"
+lemma apply_iff: "f \<in> Pi(A,B) \<Longrightarrow> <a,b>: f \<longleftrightarrow> a \<in> A \<and> f`a = b"
 apply (frule fun_is_rel)
 apply (blast intro!: apply_Pair apply_equality)
 done
@@ -110,7 +110,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)}))
-      \<longleftrightarrow>  f \<in> Pi(A,B) & (\<forall>x\<in>A. P(x, f`x))"
+      \<longleftrightarrow>  f \<in> Pi(A,B) \<and> (\<forall>x\<in>A. P(x, f`x))"
 by (blast intro: Pi_type dest: apply_type)
 
 lemma Pi_weaken_type:
@@ -126,7 +126,7 @@
 lemma range_type: "\<lbrakk><a,b> \<in> f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> b \<in> B(a)"
 by (blast dest: fun_is_rel)
 
-lemma Pair_mem_PiD: "\<lbrakk><a,b>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> a \<in> A & b \<in> B(a) & f`a = b"
+lemma Pair_mem_PiD: "\<lbrakk><a,b>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> a \<in> A \<and> b \<in> B(a) \<and> f`a = b"
 by (blast intro: domain_type range_type apply_equality)
 
 subsection\<open>Lambda Abstraction\<close>
@@ -196,7 +196,7 @@
 lemma Pi_empty2 [simp]: "(A->0) = (if A=0 then {0} else 0)"
 by (unfold Pi_def function_def, force)
 
-lemma  fun_space_empty_iff [iff]: "(A->X)=0 \<longleftrightarrow> X=0 & (A \<noteq> 0)"
+lemma  fun_space_empty_iff [iff]: "(A->X)=0 \<longleftrightarrow> X=0 \<and> (A \<noteq> 0)"
 apply auto
 apply (fast intro!: equals0I intro: lam_type)
 done
@@ -296,7 +296,7 @@
 lemma restrict_empty [simp]: "restrict(f,0) = 0"
 by (unfold restrict_def, simp)
 
-lemma restrict_iff: "z \<in> restrict(r,A) \<longleftrightarrow> z \<in> r & (\<exists>x\<in>A. \<exists>y. z = \<langle>x, y\<rangle>)"
+lemma restrict_iff: "z \<in> restrict(r,A) \<longleftrightarrow> z \<in> r \<and> (\<exists>x\<in>A. \<exists>y. z = \<langle>x, y\<rangle>)"
 by (simp add: restrict_def)
 
 lemma restrict_restrict [simp]:
--- a/src/ZF/pair.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/pair.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -18,7 +18,7 @@
 
 ML \<open>val ZF_ss = simpset_of \<^context>\<close>
 
-simproc_setup defined_Bex ("\<exists>x\<in>A. P(x) & Q(x)") = \<open>
+simproc_setup defined_Bex ("\<exists>x\<in>A. P(x) \<and> Q(x)") = \<open>
   fn _ => Quantifier1.rearrange_Bex
     (fn ctxt => unfold_tac ctxt @{thms Bex_def})
 \<close>
@@ -34,10 +34,10 @@
 lemma singleton_eq_iff [iff]: "{a} = {b} \<longleftrightarrow> a=b"
 by (rule extension [THEN iff_trans], blast)
 
-lemma doubleton_eq_iff: "{a,b} = {c,d} \<longleftrightarrow> (a=c & b=d) | (a=d & b=c)"
+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 & b=d"
+lemma Pair_iff [simp]: "<a,b> = <c,d> \<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!]
@@ -77,7 +77,7 @@
 
 text\<open>Generalizes Cartesian product\<close>
 
-lemma Sigma_iff [simp]: "<a,b>: Sigma(A,B) \<longleftrightarrow> a \<in> A & b \<in> B(a)"
+lemma Sigma_iff [simp]: "<a,b>: 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)"
--- a/src/ZF/upair.thy	Tue Sep 27 16:51:35 2022 +0100
+++ b/src/ZF/upair.thy	Tue Sep 27 17:03:23 2022 +0100
@@ -65,7 +65,7 @@
 
 subsection\<open>Rules for Binary Intersection, Defined via \<^term>\<open>Upair\<close>\<close>
 
-lemma Int_iff [simp]: "c \<in> A \<inter> B \<longleftrightarrow> (c \<in> A & c \<in> B)"
+lemma Int_iff [simp]: "c \<in> A \<inter> B \<longleftrightarrow> (c \<in> A \<and> c \<in> B)"
 apply (unfold Int_def)
 apply (blast intro: UpairI1 UpairI2 elim: UpairE)
 done
@@ -85,7 +85,7 @@
 
 subsection\<open>Rules for Set Difference, Defined via \<^term>\<open>Upair\<close>\<close>
 
-lemma Diff_iff [simp]: "c \<in> A-B \<longleftrightarrow> (c \<in> A & c\<notin>B)"
+lemma Diff_iff [simp]: "c \<in> A-B \<longleftrightarrow> (c \<in> A \<and> c\<notin>B)"
 by (unfold Diff_def, blast)
 
 lemma DiffI [intro!]: "\<lbrakk>c \<in> A;  c \<notin> B\<rbrakk> \<Longrightarrow> c \<in> A - B"
@@ -223,7 +223,7 @@
 by (unfold if_def, blast)
 
 lemma split_if [split]:
-     "P(if Q then x else y) \<longleftrightarrow> ((Q \<longrightarrow> P(x)) & (\<not>Q \<longrightarrow> P(y)))"
+     "P(if Q then x else y) \<longleftrightarrow> ((Q \<longrightarrow> P(x)) \<and> (\<not>Q \<longrightarrow> P(y)))"
 by (case_tac Q, simp_all)
 
 (** Rewrite rules for boolean case-splitting: faster than split_if [split]
@@ -238,7 +238,7 @@
 lemmas split_ifs = split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2
 
 (*Logically equivalent to split_if_mem2*)
-lemma if_iff: "a: (if P then x else y) \<longleftrightarrow> P & a \<in> x | \<not>P & a \<in> y"
+lemma if_iff: "a: (if P then x else y) \<longleftrightarrow> P \<and> a \<in> x | \<not>P \<and> a \<in> y"
 by simp
 
 lemma if_type [TC]:
@@ -247,7 +247,7 @@
 
 (** Splitting IFs in the assumptions **)
 
-lemma split_if_asm: "P(if Q then x else y) \<longleftrightarrow> (\<not>((Q & \<not>P(x)) | (\<not>Q & \<not>P(y))))"
+lemma split_if_asm: "P(if Q then x else y) \<longleftrightarrow> (\<not>((Q \<and> \<not>P(x)) | (\<not>Q \<and> \<not>P(y))))"
 by simp
 
 lemmas if_splits = split_if split_if_asm
@@ -324,19 +324,19 @@
 subsection\<open>Miniscoping of the Bounded Universal Quantifier\<close>
 
 lemma ball_simps1:
-     "(\<forall>x\<in>A. P(x) & Q)   \<longleftrightarrow> (\<forall>x\<in>A. P(x)) & (A=0 | Q)"
+     "(\<forall>x\<in>A. P(x) \<and> Q)   \<longleftrightarrow> (\<forall>x\<in>A. P(x)) \<and> (A=0 | Q)"
      "(\<forall>x\<in>A. P(x) | Q)   \<longleftrightarrow> ((\<forall>x\<in>A. P(x)) | Q)"
      "(\<forall>x\<in>A. P(x) \<longrightarrow> Q) \<longleftrightarrow> ((\<exists>x\<in>A. P(x)) \<longrightarrow> Q)"
      "(\<not>(\<forall>x\<in>A. P(x))) \<longleftrightarrow> (\<exists>x\<in>A. \<not>P(x))"
      "(\<forall>x\<in>0.P(x)) \<longleftrightarrow> True"
-     "(\<forall>x\<in>succ(i).P(x)) \<longleftrightarrow> P(i) & (\<forall>x\<in>i. P(x))"
-     "(\<forall>x\<in>cons(a,B).P(x)) \<longleftrightarrow> P(a) & (\<forall>x\<in>B. P(x))"
+     "(\<forall>x\<in>succ(i).P(x)) \<longleftrightarrow> P(i) \<and> (\<forall>x\<in>i. P(x))"
+     "(\<forall>x\<in>cons(a,B).P(x)) \<longleftrightarrow> P(a) \<and> (\<forall>x\<in>B. P(x))"
      "(\<forall>x\<in>RepFun(A,f). P(x)) \<longleftrightarrow> (\<forall>y\<in>A. P(f(y)))"
      "(\<forall>x\<in>\<Union>(A).P(x)) \<longleftrightarrow> (\<forall>y\<in>A. \<forall>x\<in>y. P(x))"
 by blast+
 
 lemma ball_simps2:
-     "(\<forall>x\<in>A. P & Q(x))   \<longleftrightarrow> (A=0 | P) & (\<forall>x\<in>A. Q(x))"
+     "(\<forall>x\<in>A. P \<and> Q(x))   \<longleftrightarrow> (A=0 | P) \<and> (\<forall>x\<in>A. Q(x))"
      "(\<forall>x\<in>A. P | Q(x))   \<longleftrightarrow> (P | (\<forall>x\<in>A. Q(x)))"
      "(\<forall>x\<in>A. P \<longrightarrow> Q(x)) \<longleftrightarrow> (P \<longrightarrow> (\<forall>x\<in>A. Q(x)))"
 by blast+
@@ -348,16 +348,16 @@
 lemmas ball_simps [simp] = ball_simps1 ball_simps2 ball_simps3
 
 lemma ball_conj_distrib:
-    "(\<forall>x\<in>A. P(x) & Q(x)) \<longleftrightarrow> ((\<forall>x\<in>A. P(x)) & (\<forall>x\<in>A. Q(x)))"
+    "(\<forall>x\<in>A. P(x) \<and> Q(x)) \<longleftrightarrow> ((\<forall>x\<in>A. P(x)) \<and> (\<forall>x\<in>A. Q(x)))"
 by blast
 
 
 subsection\<open>Miniscoping of the Bounded Existential Quantifier\<close>
 
 lemma bex_simps1:
-     "(\<exists>x\<in>A. P(x) & Q) \<longleftrightarrow> ((\<exists>x\<in>A. P(x)) & Q)"
-     "(\<exists>x\<in>A. P(x) | Q) \<longleftrightarrow> (\<exists>x\<in>A. P(x)) | (A\<noteq>0 & Q)"
-     "(\<exists>x\<in>A. P(x) \<longrightarrow> Q) \<longleftrightarrow> ((\<forall>x\<in>A. P(x)) \<longrightarrow> (A\<noteq>0 & Q))"
+     "(\<exists>x\<in>A. P(x) \<and> Q) \<longleftrightarrow> ((\<exists>x\<in>A. P(x)) \<and> Q)"
+     "(\<exists>x\<in>A. P(x) | Q) \<longleftrightarrow> (\<exists>x\<in>A. P(x)) | (A\<noteq>0 \<and> Q)"
+     "(\<exists>x\<in>A. P(x) \<longrightarrow> Q) \<longleftrightarrow> ((\<forall>x\<in>A. P(x)) \<longrightarrow> (A\<noteq>0 \<and> Q))"
      "(\<exists>x\<in>0.P(x)) \<longleftrightarrow> False"
      "(\<exists>x\<in>succ(i).P(x)) \<longleftrightarrow> P(i) | (\<exists>x\<in>i. P(x))"
      "(\<exists>x\<in>cons(a,B).P(x)) \<longleftrightarrow> P(a) | (\<exists>x\<in>B. P(x))"
@@ -367,13 +367,13 @@
 by blast+
 
 lemma bex_simps2:
-     "(\<exists>x\<in>A. P & Q(x)) \<longleftrightarrow> (P & (\<exists>x\<in>A. Q(x)))"
-     "(\<exists>x\<in>A. P | Q(x)) \<longleftrightarrow> (A\<noteq>0 & P) | (\<exists>x\<in>A. Q(x))"
+     "(\<exists>x\<in>A. P \<and> Q(x)) \<longleftrightarrow> (P \<and> (\<exists>x\<in>A. Q(x)))"
+     "(\<exists>x\<in>A. P | Q(x)) \<longleftrightarrow> (A\<noteq>0 \<and> P) | (\<exists>x\<in>A. Q(x))"
      "(\<exists>x\<in>A. P \<longrightarrow> Q(x)) \<longleftrightarrow> ((A=0 | P) \<longrightarrow> (\<exists>x\<in>A. Q(x)))"
 by blast+
 
 lemma bex_simps3:
-     "(\<exists>x\<in>Collect(A,Q).P(x)) \<longleftrightarrow> (\<exists>x\<in>A. Q(x) & P(x))"
+     "(\<exists>x\<in>Collect(A,Q).P(x)) \<longleftrightarrow> (\<exists>x\<in>A. Q(x) \<and> P(x))"
 by blast
 
 lemmas bex_simps [simp] = bex_simps1 bex_simps2 bex_simps3
@@ -391,10 +391,10 @@
 lemma bex_triv_one_point2 [simp]: "(\<exists>x\<in>A. a=x) \<longleftrightarrow> (a \<in> A)"
 by blast
 
-lemma bex_one_point1 [simp]: "(\<exists>x\<in>A. x=a & P(x)) \<longleftrightarrow> (a \<in> A & P(a))"
+lemma bex_one_point1 [simp]: "(\<exists>x\<in>A. x=a \<and> P(x)) \<longleftrightarrow> (a \<in> A \<and> P(a))"
 by blast
 
-lemma bex_one_point2 [simp]: "(\<exists>x\<in>A. a=x & P(x)) \<longleftrightarrow> (a \<in> A & P(a))"
+lemma bex_one_point2 [simp]: "(\<exists>x\<in>A. a=x \<and> P(x)) \<longleftrightarrow> (a \<in> A \<and> P(a))"
 by blast
 
 lemma ball_one_point1 [simp]: "(\<forall>x\<in>A. x=a \<longrightarrow> P(x)) \<longleftrightarrow> (a \<in> A \<longrightarrow> P(a))"