Huge tidy-up: removal of leading \!\!
authorpaulson
Mon, 13 Jul 1998 16:43:57 +0200
changeset 5137 60205b0de9b9
parent 5136 4a1ee3043101
child 5138 b02dfb930bd9
Huge tidy-up: removal of leading \!\! addsplits split_tac[split_if] now default nat_succI now included by default
src/ZF/AC.ML
src/ZF/AC/AC0_AC1.ML
src/ZF/AC/AC10_AC15.ML
src/ZF/AC/AC15_WO6.ML
src/ZF/AC/AC16_WO4.ML
src/ZF/AC/AC16_lemmas.ML
src/ZF/AC/AC17_AC1.ML
src/ZF/AC/AC18_AC19.ML
src/ZF/AC/AC1_AC17.ML
src/ZF/AC/AC2_AC6.ML
src/ZF/AC/AC7_AC9.ML
src/ZF/AC/AC_Equiv.ML
src/ZF/AC/Cardinal_aux.ML
src/ZF/AC/DC.ML
src/ZF/AC/DC_lemmas.ML
src/ZF/AC/HH.ML
src/ZF/AC/Hartog.ML
src/ZF/AC/WO1_AC.ML
src/ZF/AC/WO1_WO6.ML
src/ZF/AC/WO1_WO7.ML
src/ZF/AC/WO1_WO8.ML
src/ZF/AC/WO2_AC16.ML
src/ZF/AC/WO6_WO1.ML
src/ZF/AC/WO_AC.ML
src/ZF/AC/recfunAC16.ML
src/ZF/Arith.ML
src/ZF/Bool.ML
src/ZF/Cardinal.ML
src/ZF/CardinalArith.ML
src/ZF/Cardinal_AC.ML
src/ZF/Coind/Map.ML
src/ZF/Coind/Values.ML
src/ZF/Epsilon.ML
src/ZF/EquivClass.ML
src/ZF/Finite.ML
src/ZF/Fixedpt.ML
src/ZF/IMP/Equiv.ML
src/ZF/List.ML
src/ZF/Nat.ML
src/ZF/Order.ML
src/ZF/OrderArith.ML
src/ZF/OrderType.ML
src/ZF/Ordinal.ML
src/ZF/Perm.ML
src/ZF/QPair.ML
src/ZF/QUniv.ML
src/ZF/Rel.ML
src/ZF/Resid/Reduction.ML
src/ZF/Resid/SubUnion.ML
src/ZF/Sum.ML
src/ZF/Trancl.ML
src/ZF/Univ.ML
src/ZF/WF.ML
src/ZF/ZF.ML
src/ZF/Zorn.ML
src/ZF/ex/Acc.ML
src/ZF/ex/BT.ML
src/ZF/ex/Bin.ML
src/ZF/ex/CoUnit.ML
src/ZF/ex/Comb.ML
src/ZF/ex/Data.ML
src/ZF/ex/Integ.ML
src/ZF/ex/LList.ML
src/ZF/ex/ListN.ML
src/ZF/ex/Mutil.ML
src/ZF/ex/Ntree.ML
src/ZF/ex/Primes.ML
src/ZF/ex/Primrec.ML
src/ZF/ex/PropLog.ML
src/ZF/ex/Ramsey.ML
src/ZF/ex/Rmap.ML
src/ZF/ex/TF.ML
src/ZF/ex/Term.ML
src/ZF/func.ML
src/ZF/mono.ML
src/ZF/simpdata.ML
--- a/src/ZF/AC.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -18,7 +18,7 @@
 qed "AC_Pi";
 
 (*Using dtac, this has the advantage of DELETING the universal quantifier*)
-Goal "!!A B. ALL x:A. EX y. y:B(x) ==> EX y. y : Pi(A,B)";
+Goal "ALL x:A. EX y. y:B(x) ==> EX y. y : Pi(A,B)";
 by (rtac AC_Pi 1);
 by (etac bspec 1);
 by (assume_tac 1);
@@ -43,7 +43,7 @@
 by (ALLGOALS Blast_tac);
 qed "non_empty_family";
 
-Goal "!!A. 0 ~: A ==> EX f: A->Union(A). ALL x:A. f`x : x";
+Goal "0 ~: A ==> EX f: A->Union(A). ALL x:A. f`x : x";
 by (rtac AC_func 1);
 by (REPEAT (ares_tac [non_empty_family] 1));
 qed "AC_func0";
@@ -56,7 +56,7 @@
 by (ALLGOALS Blast_tac);
 qed "AC_func_Pow";
 
-Goal "!!A. 0 ~: A ==> EX f. f: (PROD x:A. x)";
+Goal "0 ~: A ==> EX f. f: (PROD x:A. x)";
 by (rtac AC_Pi 1);
 by (REPEAT (ares_tac [non_empty_family] 1));
 qed "AC_Pi0";
--- a/src/ZF/AC/AC0_AC1.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/AC0_AC1.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -6,19 +6,19 @@
 AC0 comes from Suppes, AC1 from Rubin & Rubin
 *)
 
-Goal "!!A. 0~:A ==> A <= Pow(Union(A))-{0}";
+Goal "0~:A ==> A <= Pow(Union(A))-{0}";
 by (Fast_tac 1);
 qed "subset_Pow_Union";
 
-Goal "!!f. [| f:(PROD X:A. X); D<=A |] ==> EX g. g:(PROD X:D. X)";
+Goal "[| f:(PROD X:A. X); D<=A |] ==> EX g. g:(PROD X:D. X)";
 by (fast_tac (claset() addSIs [restrict_type, apply_type]) 1);
 val lemma1 = result();
 
-Goalw AC_defs "!!Z. AC0 ==> AC1"; 
+Goalw AC_defs "AC0 ==> AC1"; 
 by (fast_tac (FOL_cs addSEs [lemma1, subset_Pow_Union]) 1);
 qed "AC0_AC1";
 
-Goalw AC_defs "!!Z. AC1 ==> AC0";
+Goalw AC_defs "AC1 ==> AC0";
 by (Deepen_tac 0 1);
 (*Large search space.  Faster proof by
   by (fast_tac (claset() addSIs [notI, singletonI] addSEs [notE, DiffE]) 1);
--- a/src/ZF/AC/AC10_AC15.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/AC10_AC15.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -27,13 +27,13 @@
 (*  - ex_fun_AC13_AC15                                                    *)
 (* ********************************************************************** *)
 
-Goalw [lepoll_def] "!!A. A~=0 ==> B lepoll A*B";
+Goalw [lepoll_def] "A~=0 ==> B lepoll A*B";
 by (etac not_emptyE 1);
 by (res_inst_tac [("x","lam z:B. <x,z>")] exI 1);
 by (fast_tac (claset() addSIs [snd_conv, lam_injective]) 1);
 qed "lepoll_Sigma";
 
-Goal "!!A. 0~:A ==> ALL B:{cons(0,x*nat). x:A}. ~Finite(B)";
+Goal "0~:A ==> ALL B:{cons(0,x*nat). x:A}. ~Finite(B)";
 by (rtac ballI 1);
 by (etac RepFunE 1);
 by (hyp_subst_tac 1);
@@ -44,7 +44,7 @@
 by (Fast_tac 1);
 qed "cons_times_nat_not_Finite";
 
-Goal "!!A. [| Union(C)=A; a:A |] ==> EX B:C. a:B & B <= A";
+Goal "[| Union(C)=A; a:A |] ==> EX B:C. a:B & B <= A";
 by (Fast_tac 1);
 val lemma1 = result();
 
@@ -72,7 +72,7 @@
 by (rtac lemma2 1 THEN (REPEAT (assume_tac 1)));
 val lemma3 = result();
 
-Goalw [lepoll_def] "!!A. [| A lepoll i; Ord(i) |] ==> {P(a). a:A} lepoll i";
+Goalw [lepoll_def] "[| A lepoll i; Ord(i) |] ==> {P(a). a:A} lepoll i";
 by (etac exE 1);
 by (res_inst_tac
         [("x", "lam x:RepFun(A, P). LEAST j. EX a:A. x=P(a) & f`a=j")] exI 1);
@@ -88,7 +88,7 @@
 by (fast_tac (claset() addEs [sym, left_inverse RS ssubst]) 1);
 val lemma4 = result();
 
-Goal "!!A. [| n:nat; B:A; u(B) <= cons(0, B*nat); 0:u(B); 2 lepoll u(B);  \
+Goal "[| n:nat; B:A; u(B) <= cons(0, B*nat); 0:u(B); 2 lepoll u(B);  \
 \       u(B) lepoll succ(n) |]  \
 \       ==> (lam x:A. {fst(x). x:u(x)-{0}})`B ~= 0 &  \
 \               (lam x:A. {fst(x). x:u(x)-{0}})`B <= B &  \
@@ -105,7 +105,7 @@
                 Diff_lepoll RS (nat_into_Ord RSN (2, lemma4))]) 1);
 val lemma5 = result();
 
-Goal "!!A. [| EX f. ALL B:{cons(0, x*nat). x:A}.  \
+Goal "[| EX f. ALL B:{cons(0, x*nat). x:A}.  \
 \               pairwise_disjoint(f`B) &  \
 \               sets_of_size_between(f`B, 2, succ(n)) &  \
 \               Union(f`B)=B; n:nat |]  \
@@ -123,7 +123,7 @@
 (* AC10(n) ==> AC11                                                       *)
 (* ********************************************************************** *)
 
-Goalw AC_defs "!!Z. [| n:nat; 1 le n; AC10(n) |] ==> AC11";
+Goalw AC_defs "[| n:nat; 1 le n; AC10(n) |] ==> AC11";
 by (rtac bexI 1 THEN (assume_tac 2));
 by (Fast_tac 1);
 qed "AC10_AC11";
@@ -132,7 +132,7 @@
 (* AC11 ==> AC12                                                          *)
 (* ********************************************************************** *)
 
-Goalw AC_defs "!! Z. AC11 ==> AC12";
+Goalw AC_defs "AC11 ==> AC12";
 by (fast_tac (FOL_cs addSEs [bexE] addIs [bexI]) 1);
 qed "AC11_AC12";
 
@@ -140,7 +140,7 @@
 (* AC12 ==> AC15                                                          *)
 (* ********************************************************************** *)
 
-Goalw AC_defs "!!Z. AC12 ==> AC15";
+Goalw AC_defs "AC12 ==> AC15";
 by Safe_tac;
 by (etac allE 1);
 by (etac impE 1);
@@ -166,7 +166,7 @@
 (* AC10(n) implies AC13(n) for (1 le n)                                   *)
 (* ********************************************************************** *)
 
-Goalw AC_defs "!!n. [| n:nat; 1 le n; AC10(n) |] ==> AC13(n)";
+Goalw AC_defs "[| n:nat; 1 le n; AC10(n) |] ==> AC13(n)";
 by Safe_tac;
 by (fast_tac (empty_cs addSEs [allE, cons_times_nat_not_Finite RSN (2, impE),
                                 ex_fun_AC13_AC15]) 1);
@@ -180,7 +180,7 @@
 (* AC1 ==> AC13(1)                                                        *)
 (* ********************************************************************** *)
 
-Goalw AC_defs "!!Z. AC1 ==> AC13(1)";
+Goalw AC_defs "AC1 ==> AC13(1)";
 by (rtac allI 1);
 by (etac allE 1);
 by (rtac impI 1);
@@ -197,7 +197,7 @@
 (* AC13(m) ==> AC13(n) for m <= n                                         *)
 (* ********************************************************************** *)
 
-Goalw AC_defs "!!m n. [| m:nat; n:nat; m le n; AC13(m) |] ==> AC13(n)";
+Goalw AC_defs "[| m:nat; n:nat; m le n; AC13(m) |] ==> AC13(n)";
 by (dtac nat_le_imp_lepoll 1 THEN REPEAT (assume_tac 1));
 by (fast_tac (claset() addSEs [lepoll_trans]) 1);
 qed "AC13_mono";
@@ -210,7 +210,7 @@
 (* AC13(n) ==> AC14  if 1 <= n                                            *)
 (* ********************************************************************** *)
 
-Goalw AC_defs "!!n. [| n:nat; 1 le n; AC13(n) |] ==> AC14";
+Goalw AC_defs "[| n:nat; 1 le n; AC13(n) |] ==> AC14";
 by (fast_tac (FOL_cs addIs [bexI]) 1);
 qed "AC13_AC14";
 
@@ -218,7 +218,7 @@
 (* AC14 ==> AC15                                                          *)
 (* ********************************************************************** *)
 
-Goalw AC_defs "!!Z. AC14 ==> AC15";
+Goalw AC_defs "AC14 ==> AC15";
 by (Fast_tac 1);
 qed "AC14_AC15";
 
@@ -230,11 +230,11 @@
 (* AC13(1) ==> AC1                                                        *)
 (* ********************************************************************** *)
 
-Goal "!!A. [| A~=0; A lepoll 1 |] ==> EX a. A={a}";
+Goal "[| A~=0; A lepoll 1 |] ==> EX a. A={a}";
 by (fast_tac (claset() addSEs [not_emptyE, lepoll_1_is_sing]) 1);
 qed "lemma_aux";
 
-Goal "!!f. ALL B:A. f(B)~=0 & f(B)<=B & f(B) lepoll 1  \
+Goal "ALL B:A. f(B)~=0 & f(B)<=B & f(B) lepoll 1  \
 \               ==> (lam x:A. THE y. f(x)={y}) : (PROD X:A. X)";
 by (rtac lam_type 1);
 by (dtac bspec 1 THEN (assume_tac 1));
@@ -244,7 +244,7 @@
 by (fast_tac (claset() addEs [ssubst]) 1);
 val lemma = result();
 
-Goalw AC_defs "!!Z. AC13(1) ==> AC1";
+Goalw AC_defs "AC13(1) ==> AC1";
 by (fast_tac (claset() addSEs [lemma]) 1);
 qed "AC13_AC1";
 
--- a/src/ZF/AC/AC15_WO6.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/AC15_WO6.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -7,7 +7,7 @@
 
 open AC15_WO6;
 
-Goal "!!x. Ord(x) ==> (UN a<x. F(a)) = (UN a:x. F(a))";
+Goal "Ord(x) ==> (UN a<x. F(a)) = (UN a:x. F(a))";
 by (fast_tac (claset() addSIs [ltI] addSDs [ltD]) 1);
 qed "OUN_eq_UN";
 
--- a/src/ZF/AC/AC16_WO4.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/AC16_WO4.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -11,7 +11,7 @@
 (* The case of finite set                                                 *)
 (* ********************************************************************** *)
 
-Goalw [Finite_def] "!!A. [| Finite(A); 0<m; m:nat |] ==>  \
+Goalw [Finite_def] "[| Finite(A); 0<m; m:nat |] ==>  \
 \       EX a f. Ord(a) & domain(f) = a &  \
 \               (UN b<a. f`b) = A & (ALL b<a. f`b lepoll m)";
 by (etac bexE 1);
@@ -34,7 +34,7 @@
 (* well_ord(x,r) ==> well_ord({{y,z}. y:x}, Something(x,z))  **)
 bind_thm ("well_ord_paired", (paired_bij RS bij_is_inj RS well_ord_rvimage));
 
-Goal "!!A. [| A lepoll B; ~ A lepoll C |] ==> ~ B lepoll C";
+Goal "[| A lepoll B; ~ A lepoll C |] ==> ~ B lepoll C";
 by (fast_tac (claset() addEs [notE, lepoll_trans]) 1);
 qed "lepoll_trans1";
 
@@ -43,7 +43,7 @@
 by (fast_tac (claset() addSEs [well_ord_rvimage]) 1);
 qed "well_ord_lepoll";
 
-Goal "!!X. [| well_ord(X,R); well_ord(Y,S)  \
+Goal "[| well_ord(X,R); well_ord(Y,S)  \
 \               |] ==> EX T. well_ord(X Un Y, T)";
 by (eresolve_tac [well_ord_radd RS (Un_lepoll_sum RS well_ord_lepoll)] 1);
 by (assume_tac 1);
@@ -85,7 +85,7 @@
 (* ********************************************************************** *)
 
 (*Proof simplified by LCP*)
-Goal "!!A. [| ~(EX x:A. f`x=y); f : inj(A, B); y:B |]  \
+Goal "[| ~(EX x:A. f`x=y); f : inj(A, B); y:B |]  \
 \       ==> (lam a:succ(A). if(a=A, y, f`a)) : inj(succ(A), B)";
 by (res_inst_tac [("d","%z. if(z=y, A, converse(f)`z)")] lam_injective 1);
 by (ALLGOALS
@@ -125,7 +125,7 @@
 val ordertype_eqpoll =
         ordermap_bij RS (exI RS (eqpoll_def RS def_imp_iff RS iffD2));
 
-Goal "!!y. [| well_ord(y,R); ~Finite(y); n:nat  \
+Goal "[| well_ord(y,R); ~Finite(y); n:nat  \
 \       |] ==> EX z. z<=y & n eqpoll z";
 by (etac nat_lepoll_imp_ex_eqpoll_n 1);
 by (resolve_tac [ordertype_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll
@@ -135,14 +135,14 @@
                         RS lepoll_infinite]) 1);
 qed "ex_subset_eqpoll_n";
 
-Goalw [lesspoll_def] "!!n. n: nat ==> n lesspoll nat";
+Goalw [lesspoll_def] "n: nat ==> n lesspoll nat";
 by (fast_tac (claset() addSEs [Ord_nat RSN (2, ltI) RS leI RS le_imp_lepoll,
         eqpoll_sym RS eqpoll_imp_lepoll]
         addIs [Ord_nat RSN (2, nat_succI RS ltI) RS leI
         RS le_imp_lepoll RS lepoll_trans RS succ_lepoll_natE]) 1);
 qed "n_lesspoll_nat";
 
-Goal "!!y. [| well_ord(y,R); ~Finite(y); k eqpoll a; a <= y; k: nat |]  \
+Goal "[| well_ord(y,R); ~Finite(y); k eqpoll a; a <= y; k: nat |]  \
 \       ==> y - a eqpoll y";
 by (fast_tac (empty_cs addIs [lepoll_lesspoll_lesspoll]
         addSIs [Card_cardinal, Diff_lesspoll_eqpoll_Card RS eqpoll_trans,
@@ -156,11 +156,11 @@
                 RS lepoll_infinite]) 1);
 qed "Diff_Finite_eqpoll";
 
-Goal "!!x. [| a<=y; b:y-a; u:x |] ==> cons(b, cons(u, a)) : Pow(x Un y)";
+Goal "[| a<=y; b:y-a; u:x |] ==> cons(b, cons(u, a)) : Pow(x Un y)";
 by (Fast_tac 1);
 qed "cons_cons_subset";
 
-Goal "!!x. [| a eqpoll k; a<=y; b:y-a; u:x; x Int y = 0  \
+Goal "[| a eqpoll k; a<=y; b:y-a; u:x; x Int y = 0  \
 \       |] ==> cons(b, cons(u, a)) eqpoll succ(succ(k))";
 by (fast_tac (claset() addSIs [cons_eqpoll_succ] addEs [equals0D]) 1);
 qed "cons_cons_eqpoll";
@@ -176,7 +176,7 @@
                 addSEs [subset_imp_lepoll RSN (2, lepoll_trans)]) 1);
 qed "s_uI";
 
-Goalw [s_u_def] "!!v. v : s_u(u, t_n, k, y) ==> u : v";
+Goalw [s_u_def] "v : s_u(u, t_n, k, y) ==> u : v";
 by (Fast_tac 1);
 qed "in_s_u_imp_u_in";
 
@@ -213,7 +213,8 @@
 by (Fast_tac 1);
 by (dtac cons_eqpoll_succ 1);
 by (Fast_tac 1);
-by (fast_tac (claset() addSIs [nat_succI]
+by (fast_tac 
+    (claset() 
         addSEs [[eqpoll_sym RS eqpoll_imp_lepoll, subset_imp_lepoll] MRS
         (lepoll_trans RS lepoll_trans) RS succ_lepoll_natE]) 1);
 qed "set_eq_cons";
@@ -230,15 +231,15 @@
 by (REPEAT (Fast_tac 1));
 qed "the_eq_cons";
 
-Goal "!!a. [| cons(x,a) = cons(y,a); x~: a |] ==> x = y ";
+Goal "[| cons(x,a) = cons(y,a); x~: a |] ==> x = y ";
 by (fast_tac (claset() addSEs [equalityE]) 1);
 qed "cons_eqE";
 
-Goal "!!A. A = B ==> A Int C = B Int C";
+Goal "A = B ==> A Int C = B Int C";
 by (Asm_simp_tac 1);
 qed "eq_imp_Int_eq";
 
-Goal "!!a. [| a=b; a=c; b=d |] ==> c=d";
+Goal "[| a=b; a=c; b=d |] ==> c=d";
 by (Asm_full_simp_tac 1);
 qed "msubst";
 
@@ -300,7 +301,7 @@
 by (Fast_tac 1);
 qed "eqpoll_sum_imp_Diff_lepoll_lemma";
 
-Goal "!!k. [| A eqpoll succ(k #+ m); B<=A; succ(k) lepoll B;  \
+Goal "[| A eqpoll succ(k #+ m); B<=A; succ(k) lepoll B;  \
 \       k:nat; m:nat |]  \
 \       ==> A-B lepoll m";
 by (dresolve_tac [add_succ RS ssubst] 1);
@@ -332,7 +333,7 @@
 by (Fast_tac 1);
 qed "eqpoll_sum_imp_Diff_eqpoll_lemma";
 
-Goal "!!k. [| A eqpoll succ(k #+ m); B<=A; succ(k) eqpoll B;  \
+Goal "[| A eqpoll succ(k #+ m); B<=A; succ(k) eqpoll B;  \
 \       k:nat; m:nat |]  \
 \       ==> A-B eqpoll m";
 by (dresolve_tac [add_succ RS ssubst] 1);
@@ -345,12 +346,12 @@
 (* back to the second part                                                *)
 (* ********************************************************************** *)
 
-Goal "!!w. [| x Int y = 0; w <= x Un y |]  \
+Goal "[| x Int y = 0; w <= x Un y |]  \
 \        ==> w Int (x - {u}) = w - cons(u, w Int y)";
 by (fast_tac (claset() addEs [equals0D]) 1);
 qed "w_Int_eq_w_Diff";
 
-Goal "!!w. [| w:{v:s_u(u, t_n, succ(l), y). a <= v};  \
+Goal "[| w:{v:s_u(u, t_n, succ(l), y). a <= v};  \
 \       l eqpoll a; t_n <= {v:Pow(x Un y). v eqpoll succ(succ(l) #+ m)};  \
 \       m:nat; l:nat; x Int y = 0; u : x;  \
 \       ALL v:s_u(u, t_n, succ(l), y). succ(l) eqpoll v Int y  \
@@ -361,10 +362,10 @@
 by (fast_tac (claset() addEs [equals0D] addSDs [bspec]
         addDs [s_u_subset RS subsetD]
         addSEs [eqpoll_sym RS cons_eqpoll_succ RS eqpoll_sym, in_s_u_imp_u_in]
-        addSIs [nat_succI, eqpoll_sum_imp_Diff_eqpoll]) 1);
+        addSIs [eqpoll_sum_imp_Diff_eqpoll]) 1);
 qed "w_Int_eqpoll_m";
 
-Goal "!!m. [| 0<m; x eqpoll m; m:nat |] ==> x ~= 0";
+Goal "[| 0<m; x eqpoll m; m:nat |] ==> x ~= 0";
 by (fast_tac (empty_cs
         addSEs [mem_irrefl, ltE, eqpoll_succ_imp_not_empty, natE]) 1);
 qed "eqpoll_m_not_empty";
@@ -410,7 +411,7 @@
 by (fast_tac (claset() addSEs [s_u_subset RS subsetD, in_s_u_imp_u_in]) 1);
 qed "subset_s_u_lepoll_w";
 
-Goal "!!k. [| 0<k; k:nat |] ==> EX l:nat. k = succ(l)";
+Goal "[| 0<k; k:nat |] ==> EX l:nat. k = succ(l)";
 by (etac natE 1);
 by (fast_tac (empty_cs addSEs [ltE, mem_irrefl]) 1);
 by (fast_tac (empty_cs addSIs [refl, bexI]) 1);
@@ -443,7 +444,7 @@
                       addSIs [lepoll_refl]) 1);
 qed "subsets_lepoll_0_eq_unit";
 
-Goal "!!X. [| well_ord(X, R); ~Finite(X); n:nat |]  \
+Goal "[| well_ord(X, R); ~Finite(X); n:nat |]  \
 \               ==> EX S. well_ord({Y: Pow(X) . Y eqpoll succ(n)}, S)";
 by (resolve_tac [well_ord_infinite_subsets_eqpoll_X
         RS (eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1
@@ -451,14 +452,14 @@
 by (fast_tac (claset() addSEs [bij_is_inj RS well_ord_rvimage]) 1);
 qed "well_ord_subsets_eqpoll_n";
 
-Goal "!!n. n:nat ==> {z:Pow(y). z lepoll succ(n)} =  \
+Goal "n:nat ==> {z:Pow(y). z lepoll succ(n)} =  \
 \       {z:Pow(y). z lepoll n} Un {z:Pow(y). z eqpoll succ(n)}";
 by (fast_tac (claset() addIs [le_refl, leI, le_imp_lepoll]
                 addSDs [lepoll_succ_disj]
                 addSEs [nat_into_Ord, lepoll_trans, eqpoll_imp_lepoll]) 1);
 qed "subsets_lepoll_succ";
 
-Goal "!!n. n:nat ==>  \
+Goal "n:nat ==>  \
 \       {z:Pow(y). z lepoll n} Int {z:Pow(y). z eqpoll succ(n)} = 0";
 by (fast_tac (claset() addSEs [eqpoll_sym RS eqpoll_imp_lepoll 
                 RS lepoll_trans RS succ_lepoll_natE]
@@ -486,7 +487,7 @@
 (* well_ord_subsets_lepoll_n                                              *)
 (* ********************************************************************** *)
 
-Goal "!!y r. [| well_ord(y,r); ~Finite(y); n:nat |] ==>  \
+Goal "[| well_ord(y,r); ~Finite(y); n:nat |] ==>  \
 \       EX R. well_ord({z:Pow(y). z lepoll n}, R)";
 by (etac nat_induct 1);
 by (fast_tac (claset() addSIs [well_ord_unit]
@@ -509,7 +510,7 @@
                 RSN (2, lepoll_trans))]) 1);
 qed "LL_subset";
 
-Goal "!!x. [| t_n <= {v:Pow(x Un y). v eqpoll n}; \
+Goal "[| t_n <= {v:Pow(x Un y). v eqpoll n}; \
 \               well_ord(y, R); ~Finite(y); n:nat  \
 \               |] ==> EX S. well_ord(LL(t_n, k, y), S)";
 by (fast_tac (FOL_cs addIs [exI]
@@ -559,7 +560,7 @@
 by (Fast_tac 1);
 qed "exists_in_MM";
 
-Goalw [LL_def] "!!w. w : MM(t_n, k, y) ==> w Int y : LL(t_n, k, y)";
+Goalw [LL_def] "w : MM(t_n, k, y) ==> w Int y : LL(t_n, k, y)";
 by (Fast_tac 1);
 qed "Int_in_LL";
 
@@ -695,7 +696,7 @@
 by (forward_tac [infinite_Un] 1 THEN (mp_tac 1));
 by (REPEAT (eresolve_tac [exE, conjE] 1));
 by (resolve_tac [well_ord_LL RS exE] 1 THEN REPEAT (assume_tac 1));
-by (fast_tac (claset() addSIs [nat_succI, add_type]) 1);
+by (fast_tac (claset() addSIs [add_type]) 1);
 by (res_inst_tac [("x","ordertype(LL(T, succ(k), y), x)")] exI 1);
 by (res_inst_tac [("x","lam b:ordertype(LL(T, succ(k), y), x).  \
 \       (GG(T, succ(k), y)) `  \
--- a/src/ZF/AC/AC16_lemmas.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/AC16_lemmas.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -7,7 +7,7 @@
 
 open AC16_lemmas;
 
-Goal "!!a. a~:A ==> cons(a,A)-{a}=A";
+Goal "a~:A ==> cons(a,A)-{a}=A";
 by (Fast_tac 1);
 qed "cons_Diff_eq";
 
@@ -61,7 +61,7 @@
 by (resolve_tac [eqpoll_RepFun_sing RS eqpoll_sym] 1);
 qed "subsets_eqpoll_1_eqpoll";
 
-Goal "!!x. [| InfCard(x); y:Pow(x); y eqpoll succ(z) |]  \
+Goal "[| InfCard(x); y:Pow(x); y eqpoll succ(z) |]  \
 \               ==> (LEAST i. i:y) : y";
 by (eresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS 
                 succ_lepoll_imp_not_empty RS not_emptyE] 1);
@@ -70,7 +70,7 @@
         addEs [Ord_in_Ord]) 1);
 qed "InfCard_Least_in";
 
-Goalw [lepoll_def] "!!i. [| InfCard(x); n:nat |] ==>  \
+Goalw [lepoll_def] "[| InfCard(x); n:nat |] ==>  \
 \       {y:Pow(x). y eqpoll succ(succ(n))} lepoll  \
 \       x*{y:Pow(x). y eqpoll succ(n)}";
 by (res_inst_tac [("x","lam y:{y:Pow(x). y eqpoll succ(succ(n))}. \
@@ -104,7 +104,7 @@
 by (fast_tac (claset() addSEs [leE,ltE]) 1);
 qed "set_of_Ord_succ_Union";
 
-Goal "!!i. j<=i ==> i ~: j";
+Goal "j<=i ==> i ~: j";
 by (fast_tac (claset() addSEs [mem_irrefl]) 1);
 qed "subset_not_mem";
 
@@ -117,15 +117,15 @@
 by (Fast_tac 1);
 qed "Union_cons_eq_succ_Union";
 
-Goal "!!i. [| Ord(i); Ord(j) |] ==> i Un j = i | i Un j = j";
+Goal "[| Ord(i); Ord(j) |] ==> i Un j = i | i Un j = j";
 by (fast_tac (claset() addSDs [le_imp_subset] addEs [Ord_linear_le]) 1);
 qed "Un_Ord_disj";
 
-Goal "!!X. x:X ==> Union(X) = x Un Union(X-{x})";
+Goal "x:X ==> Union(X) = x Un Union(X-{x})";
 by (Fast_tac 1);
 qed "Union_eq_Un";
 
-Goal "!!n. n:nat ==>  \
+Goal "n:nat ==>  \
 \       ALL z. (ALL y:z. Ord(y)) & z eqpoll n & z~=0 --> Union(z) : z";
 by (etac nat_induct 1);
 by (fast_tac (claset() addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]) 1);
@@ -149,13 +149,13 @@
 by (rtac subst_elem 1 THEN (TRYALL assume_tac));
 qed "Union_in_lemma";
 
-Goal "!!z. [| ALL x:z. Ord(x); z eqpoll n; z~=0; n:nat |]  \
+Goal "[| ALL x:z. Ord(x); z eqpoll n; z~=0; n:nat |]  \
 \               ==> Union(z) : z";
 by (dtac Union_in_lemma 1);
 by (fast_tac FOL_cs 1);
 qed "Union_in";
 
-Goal "!!x. [| InfCard(x); z:Pow(x); z eqpoll n; n:nat |]  \
+Goal "[| InfCard(x); z:Pow(x); z eqpoll n; n:nat |]  \
 \               ==> succ(Union(z)) : x";
 by (resolve_tac [Limit_has_succ RS ltE] 1 THEN (assume_tac 3));
 by (etac InfCard_is_Limit 1);
@@ -170,7 +170,7 @@
 		 (2, InfCard_is_Card RS Card_is_Ord RS Ord_in_Ord)]) 1);
 qed "succ_Union_in_x";
 
-Goalw [lepoll_def] "!!X. [| InfCard(x); n:nat |] ==>  \
+Goalw [lepoll_def] "[| InfCard(x); n:nat |] ==>  \
 \       {y:Pow(x). y eqpoll succ(n)} lepoll  \
 \       {y:Pow(x). y eqpoll succ(succ(n))}";
 by (res_inst_tac [("x","lam z:{y:Pow(x). y eqpoll succ(n)}.  \
@@ -185,10 +185,10 @@
                     addSIs [succ_Union_not_mem] 
                     addSDs [InfCard_is_Card RS Card_is_Ord] 
                     addEs  [Ord_in_Ord]) 2);
-by (fast_tac (claset() addSIs [succ_Union_in_x, nat_succI]) 1);
+by (fast_tac (claset() addSIs [succ_Union_in_x]) 1);
 qed "succ_lepoll_succ_succ";
 
-Goal "!!X. [| InfCard(X); n:nat |]  \
+Goal "[| InfCard(X); n:nat |]  \
 \       ==> {Y:Pow(X). Y eqpoll succ(n)} eqpoll X";
 by (etac nat_induct 1);
 by (rtac subsets_eqpoll_1_eqpoll 1);
@@ -205,18 +205,18 @@
         addSIs [succ_lepoll_succ_succ]) 1);
 qed "subsets_eqpoll_X";
 
-Goalw [surj_def] "!!f. [| f:surj(A,B); y<=B |]  \
+Goalw [surj_def] "[| f:surj(A,B); y<=B |]  \
 \       ==> f``(converse(f)``y) = y";
 by (fast_tac (claset() addDs [apply_equality2]
 	              addEs [apply_iff RS iffD2]) 1);
 qed "image_vimage_eq";
 
-Goal "!!f. [| f:inj(A,B); y<=A |] ==> converse(f)``(f``y) = y";
+Goal "[| f:inj(A,B); y<=A |] ==> converse(f)``(f``y) = y";
 by (fast_tac (claset() addSEs [inj_is_fun RS apply_Pair]
                 addDs [inj_equality]) 1);
 qed "vimage_image_eq";
 
-Goalw [eqpoll_def] "!!A B. A eqpoll B  \
+Goalw [eqpoll_def] "A eqpoll B  \
 \       ==> {Y:Pow(A). Y eqpoll n} eqpoll {Y:Pow(B). Y eqpoll n}";
 by (etac exE 1);
 by (res_inst_tac [("x","lam X:{Y:Pow(A). EX f. f : bij(Y, n)}. f``X")] exI 1);
@@ -232,22 +232,22 @@
 by (fast_tac (claset() addSEs [bij_is_surj RS image_vimage_eq]) 1);
 qed "subsets_eqpoll";
 
-Goalw [WO2_def] "!!X. WO2 ==> EX a. Card(a) & X eqpoll a";
+Goalw [WO2_def] "WO2 ==> EX a. Card(a) & X eqpoll a";
 by (REPEAT (eresolve_tac [allE,exE,conjE] 1));
 by (fast_tac (claset() addSEs [well_ord_Memrel RS well_ord_cardinal_eqpoll RS
                 (eqpoll_sym RSN (2, eqpoll_trans)) RS eqpoll_sym]
                 addSIs [Card_cardinal]) 1);
 qed "WO2_imp_ex_Card";
 
-Goal "!!X. [| X lepoll Y; ~Finite(X) |] ==> ~Finite(Y)";
+Goal "[| X lepoll Y; ~Finite(X) |] ==> ~Finite(Y)";
 by (fast_tac (empty_cs addEs [notE, lepoll_Finite] addSIs [notI]) 1); 
 qed "lepoll_infinite";
 
-Goalw [InfCard_def] "!!X. [| ~Finite(X); Card(X) |] ==> InfCard(X)";
+Goalw [InfCard_def] "[| ~Finite(X); Card(X) |] ==> InfCard(X)";
 by (fast_tac (claset() addSEs [Card_is_Ord RS nat_le_infinite_Ord]) 1);
 qed "infinite_Card_is_InfCard";
 
-Goal "!!X n. [| WO2; n:nat; ~Finite(X) |]  \
+Goal "[| WO2; n:nat; ~Finite(X) |]  \
 \       ==> {Y:Pow(X). Y eqpoll succ(n)} eqpoll X";
 by (dtac WO2_imp_ex_Card 1);
 by (REPEAT (eresolve_tac [allE,exE,conjE] 1));
@@ -259,12 +259,12 @@
 by (etac eqpoll_sym 1);
 qed "WO2_infinite_subsets_eqpoll_X";
 
-Goal "!!X. well_ord(X,R) ==> EX a. Card(a) & X eqpoll a";
+Goal "well_ord(X,R) ==> EX a. Card(a) & X eqpoll a";
 by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll RS eqpoll_sym]
                 addSIs [Card_cardinal]) 1);
 qed "well_ord_imp_ex_Card";
 
-Goal "!!X n. [| well_ord(X,R); n:nat; ~Finite(X) |]  \
+Goal "[| well_ord(X,R); n:nat; ~Finite(X) |]  \
 \               ==> {Y:Pow(X). Y eqpoll succ(n)} eqpoll X";
 by (dtac well_ord_imp_ex_Card 1);
 by (REPEAT (eresolve_tac [allE,exE,conjE] 1));
--- a/src/ZF/AC/AC17_AC1.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/AC17_AC1.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -11,7 +11,7 @@
 (* more properties of HH                                                   *)
 (* *********************************************************************** *)
 
-Goal "!!f. [| x - (UN j:LEAST i. HH(lam X:Pow(x)-{0}. {f`X}, x, i) = {x}. \
+Goal "[| x - (UN j:LEAST i. HH(lam X:Pow(x)-{0}. {f`X}, x, i) = {x}. \
 \       HH(lam X:Pow(x)-{0}. {f`X}, x, j)) = 0;  \
 \       f : Pow(x)-{0} -> x |]  \
 \       ==> EX r. well_ord(x,r)";
@@ -25,7 +25,7 @@
 (* theorems closer to the proof                                            *)
 (* *********************************************************************** *)
 
-Goalw AC_defs "!!Z. ~AC1 ==>  \
+Goalw AC_defs "~AC1 ==>  \
 \               EX A. ALL f:Pow(A)-{0} -> A. EX u:Pow(A)-{0}. f`u ~: u";
 by (etac swap 1);
 by (rtac allI 1);
@@ -37,7 +37,7 @@
 by (fast_tac (claset() addSIs [restrict_type]) 1);
 qed "not_AC1_imp_ex";
 
-Goal "!!x. [| ALL f:Pow(x) - {0} -> x. EX u: Pow(x) - {0}. f`u~:u;  \
+Goal "[| ALL f:Pow(x) - {0} -> x. EX u: Pow(x) - {0}. f`u~:u;  \
 \       EX f: Pow(x)-{0}->x. \
 \       x - (UN a:(LEAST i. HH(lam X:Pow(x)-{0}. {f`X},x,i)={x}).  \
 \       HH(lam X:Pow(x)-{0}. {f`X},x,a)) = 0 |] \
@@ -53,24 +53,24 @@
 by (Fast_tac 1);
 val lemma1 = result();
 
-Goal "!!x. ~ (EX f: Pow(x)-{0}->x. x - F(f) = 0)  \
+Goal "~ (EX f: Pow(x)-{0}->x. x - F(f) = 0)  \
 \       ==> (lam f: Pow(x)-{0}->x. x - F(f))  \
 \               : (Pow(x) -{0} -> x) -> Pow(x) - {0}";
 by (fast_tac (claset() addSIs [lam_type] addSDs [Diff_eq_0_iff RS iffD1]) 1);
 val lemma2 = result();
 
-Goal "!!f. [| f`Z : Z; Z:Pow(x)-{0} |] ==>  \
+Goal "[| f`Z : Z; Z:Pow(x)-{0} |] ==>  \
 \       (lam X:Pow(x)-{0}. {f`X})`Z : Pow(Z)-{0}";
 by (Asm_full_simp_tac 1);
 by (fast_tac (claset() addSDs [equals0D]) 1);
 val lemma3 = result();
 
-Goal "!!z. EX f:F. f`((lam f:F. Q(f))`f) : (lam f:F. Q(f))`f  \
+Goal "EX f:F. f`((lam f:F. Q(f))`f) : (lam f:F. Q(f))`f  \
 \       ==> EX f:F. f`Q(f) : Q(f)";
 by (Asm_full_simp_tac 1);
 val lemma4 = result();
 
-Goalw [AC17_def] "!!Z. AC17 ==> AC1";
+Goalw [AC17_def] "AC17 ==> AC1";
 by (rtac classical 1);
 by (eresolve_tac [not_AC1_imp_ex RS exE] 1);
 by (excluded_middle_tac
--- a/src/ZF/AC/AC18_AC19.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/AC18_AC19.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -11,7 +11,7 @@
 (* AC1 ==> AC18                                                           *)
 (* ********************************************************************** *)
 
-Goal "!!f. [| f: (PROD b:{P(a). a:A}. b); ALL a:A. P(a)<=Q(a) |]  \
+Goal "[| f: (PROD b:{P(a). a:A}. b); ALL a:A. P(a)<=Q(a) |]  \
 \               ==> (lam a:A. f`P(a)):(PROD a:A. Q(a))";
 by (rtac lam_type 1);
 by (dtac apply_type 1);
@@ -20,7 +20,7 @@
 by (etac subsetD 1 THEN (assume_tac 1));
 qed "PROD_subsets";
 
-Goal "!!X. [| ALL A. 0 ~: A --> (EX f. f : (PROD X:A. X)); A ~= 0 |] ==>  \
+Goal "[| ALL A. 0 ~: A --> (EX f. f : (PROD X:A. X)); A ~= 0 |] ==>  \
 \  (INT a:A. UN b:B(a). X(a, b)) <= (UN f:PROD a:A. B(a). INT a:A. X(a, f`a))";
 by (rtac subsetI 1);
 by (eres_inst_tac [("x","{{b:B(a). x:X(a,b)}. a:A}")] allE 1);
@@ -62,7 +62,7 @@
                 addSDs [sym RS (RepFun_eq_0_iff RS iffD1)]) 1);
 qed "RepRep_conj";
 
-Goal "!!c. [|c : a; x = c Un {0}; x ~: a |] ==> x - {0} : a";
+Goal "[|c : a; x = c Un {0}; x ~: a |] ==> x - {0} : a";
 by (hyp_subst_tac 1);
 by (rtac subst_elem 1 THEN (assume_tac 1));
 by (rtac equalityI 1);
@@ -80,7 +80,7 @@
                 addSDs [apply_type]) 1);
 val lemma1_2 = result();
 
-Goal  "!!A. EX f. f: (PROD B:{u_(a). a:A}. B) ==> EX f. f: (PROD B:A. B)";
+Goal  "EX f. f: (PROD B:{u_(a). a:A}. B) ==> EX f. f: (PROD B:A. B)";
 by (etac exE 1);
 by (res_inst_tac
         [("x","lam a:A. if(f`(u_(a)) : a, f`(u_(a)), f`(u_(a))-{0})")] exI 1);
@@ -91,21 +91,21 @@
 by (fast_tac (claset() addSEs [lemma1_2]) 1);
 val lemma1 = result();
 
-Goalw [u_def] "!!a. a~=0 ==> 0: (UN b:u_(a). b)";
+Goalw [u_def] "a~=0 ==> 0: (UN b:u_(a). b)";
 by (fast_tac (claset() addSEs [not_emptyE] addSIs [UN_I, RepFunI]) 1);
 val lemma2_1 = result();
 
-Goal "!!A C. [| A~=0; 0~:A |] ==> (INT x:{u_(a). a:A}. UN b:x. b) ~= 0";
+Goal "[| A~=0; 0~:A |] ==> (INT x:{u_(a). a:A}. UN b:x. b) ~= 0";
 by (etac not_emptyE 1);
 by (res_inst_tac [("a","0")] not_emptyI 1);
 by (fast_tac (claset() addSIs [INT_I, RepFunI, lemma2_1] addSEs [RepFunE]) 1);
 val lemma2 = result();
 
-Goal "!!F. (UN f:F. P(f)) ~= 0 ==> F ~= 0";
+Goal "(UN f:F. P(f)) ~= 0 ==> F ~= 0";
 by (fast_tac (claset() addSEs [not_emptyE]) 1);
 val lemma3 = result();
 
-Goalw AC_defs "!!Z. AC19 ==> AC1";
+Goalw AC_defs "AC19 ==> AC1";
 by (REPEAT (resolve_tac [allI,impI] 1));
 by (excluded_middle_tac "A=0" 1);
 by (fast_tac (claset() addSIs [exI, empty_fun]) 2);
--- a/src/ZF/AC/AC1_AC17.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/AC1_AC17.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -5,13 +5,13 @@
 The proof of AC1 ==> AC17
 *)
 
-Goal "!!f. f : (PROD X:Pow(A) - {0}. X) ==> f : (Pow(A) - {0} -> A)";
+Goal "f : (PROD X:Pow(A) - {0}. X) ==> f : (Pow(A) - {0} -> A)";
 by (rtac Pi_type 1 THEN (assume_tac 1));
 by (dtac apply_type 1 THEN (assume_tac 1));
 by (Fast_tac 1);
 val lemma1 = result();
 
-Goalw AC_defs "!!Z. AC1 ==> AC17";
+Goalw AC_defs "AC1 ==> AC17";
 by (rtac allI 1);
 by (rtac ballI 1);
 by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1);
--- a/src/ZF/AC/AC2_AC6.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/AC2_AC6.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -14,7 +14,7 @@
 (* AC1 ==> AC2                                                            *)
 (* ********************************************************************** *)
 
-Goal "!!B. [| B:A; f:(PROD X:A. X); 0~:A |]  \
+Goal "[| B:A; f:(PROD X:A. X); 0~:A |]  \
 \               ==> {f`B} <= B Int {f`C. C:A}";
 by (fast_tac (claset() addSEs [apply_type]) 1);
 val lemma1 = result();
@@ -24,7 +24,7 @@
 by (fast_tac (claset() addSEs [equals0D]) 1);
 val lemma2 = result();
 
-Goalw AC_defs "!!Z. AC1 ==> AC2"; 
+Goalw AC_defs "AC1 ==> AC2"; 
 by (rtac allI 1);
 by (rtac impI 1);
 by (REPEAT (eresolve_tac [asm_rl,conjE,allE,exE,impE] 1));
@@ -38,12 +38,12 @@
 (* AC2 ==> AC1                                                            *)
 (* ********************************************************************** *)
 
-Goal "!!A. 0~:A ==> 0 ~: {B*{B}. B:A}";
+Goal "0~:A ==> 0 ~: {B*{B}. B:A}";
 by (fast_tac (claset() addSDs [sym RS (Sigma_empty_iff RS iffD1)]
         addSEs [RepFunE, equals0D]) 1);
 val lemma1 = result();
 
-Goal "!!A. [| X*{X} Int C = {y}; X:A |]  \
+Goal "[| X*{X} Int C = {y}; X:A |]  \
 \               ==> (THE y. X*{X} Int C = {y}): X*A";
 by (rtac subst_elem 1);
 by (fast_tac (claset() addSIs [the_equality]
@@ -51,7 +51,7 @@
 by (fast_tac (claset() addSEs [equalityE, make_elim singleton_subsetD]) 1);
 val lemma2 = result();
 
-Goal "!!A. ALL D:{E*{E}. E:A}. EX y. D Int C = {y}  \
+Goal "ALL D:{E*{E}. E:A}. EX y. D Int C = {y}  \
 \       ==> (lam x:A. fst(THE z. (x*{x} Int C = {z}))) :  \
 \               (PROD X:A. X) ";
 by (fast_tac (claset() addSEs [lemma2]
@@ -71,11 +71,11 @@
 (* AC1 ==> AC4                                                            *)
 (* ********************************************************************** *)
 
-Goal "!!R. 0 ~: {R``{x}. x:domain(R)}";
+Goal "0 ~: {R``{x}. x:domain(R)}";
 by (fast_tac (claset() addEs [sym RS equals0D]) 1);
 val lemma = result();
 
-Goalw AC_defs "!!Z. AC1 ==> AC4";
+Goalw AC_defs "AC1 ==> AC4";
 by (REPEAT (resolve_tac [allI, impI] 1));
 by (REPEAT (eresolve_tac [allE, lemma RSN (2, impE), exE] 1));
 by (best_tac (claset() addSIs [lam_type] addSEs [apply_type]) 1);
@@ -86,19 +86,19 @@
 (* AC4 ==> AC3                                                            *)
 (* ********************************************************************** *)
 
-Goal "!!f. f:A->B ==> (UN z:A. {z}*f`z) <= A*Union(B)";
+Goal "f:A->B ==> (UN z:A. {z}*f`z) <= A*Union(B)";
 by (fast_tac (claset() addSDs [apply_type]) 1);
 val lemma1 = result();
 
-Goal "!!f. domain(UN z:A. {z}*f(z)) = {a:A. f(a)~=0}";
+Goal "domain(UN z:A. {z}*f(z)) = {a:A. f(a)~=0}";
 by (fast_tac (claset() addSIs [not_emptyI] addDs [range_type]) 1);
 val lemma2 = result();
 
-Goal "!!f. x:A ==> (UN z:A. {z}*f(z))``{x} = f(x)";
+Goal "x:A ==> (UN z:A. {z}*f(z))``{x} = f(x)";
 by (Fast_tac 1);
 val lemma3 = result();
 
-Goalw AC_defs "!!Z. AC4 ==> AC3";
+Goalw AC_defs "AC4 ==> AC3";
 by (REPEAT (resolve_tac [allI,ballI] 1));
 by (REPEAT (eresolve_tac [allE,impE] 1));
 by (etac lemma1 1);
@@ -110,13 +110,13 @@
 (* AC3 ==> AC1                                                            *)
 (* ********************************************************************** *)
 
-Goal "!!A. b~:A ==> (PROD x:{a:A. id(A)`a~=b}. id(A)`x) = (PROD x:A. x)";
+Goal "b~:A ==> (PROD x:{a:A. id(A)`a~=b}. id(A)`x) = (PROD x:A. x)";
 by (asm_full_simp_tac (simpset() addsimps [id_def] addcongs [Pi_cong]) 1);
 by (res_inst_tac [("b","A")] subst_context 1);
 by (Fast_tac 1);
 val lemma = result();
 
-Goalw AC_defs "!!Z. AC3 ==> AC1";
+Goalw AC_defs "AC3 ==> AC1";
 by (REPEAT (resolve_tac [allI, impI] 1));
 by (REPEAT (eresolve_tac [allE, ballE] 1));
 by (fast_tac (claset() addSIs [id_type]) 2);
@@ -148,11 +148,11 @@
 (* AC5 ==> AC4, Rubin & Rubin, p. 11                                      *)
 (* ********************************************************************** *)
 
-Goal "!!A. R <= A*B ==> (lam x:R. fst(x)) : R -> A";
+Goal "R <= A*B ==> (lam x:R. fst(x)) : R -> A";
 by (fast_tac (claset() addSIs [lam_type, fst_type]) 1);
 val lemma1 = result();
 
-Goalw [range_def] "!!A. R <= A*B ==> range(lam x:R. fst(x)) = domain(R)";
+Goalw [range_def] "R <= A*B ==> range(lam x:R. fst(x)) = domain(R)";
 by (rtac equalityI 1);
 by (fast_tac (claset() addSEs [lamE]
                 addEs [subst_elem]
@@ -163,13 +163,13 @@
 by (fast_tac (claset() addSEs [lamI RS subst_elem] addIs [fst_conv RS ssubst]) 1);
 val lemma2 = result();
 
-Goal "!!A. [| EX f: A->C. P(f,domain(f)); A=B |] ==>  EX f: B->C. P(f,B)";
+Goal "[| EX f: A->C. P(f,domain(f)); A=B |] ==>  EX f: B->C. P(f,B)";
 by (etac bexE 1);
 by (forward_tac [domain_of_fun] 1);
 by (Fast_tac 1);
 val lemma3 = result();
 
-Goal "!!g. [| R <= A*B; g: C->R; ALL x:C. (lam z:R. fst(z))` (g`x) = x |] \
+Goal "[| R <= A*B; g: C->R; ALL x:C. (lam z:R. fst(z))` (g`x) = x |] \
 \               ==> (lam x:C. snd(g`x)): (PROD x:C. R``{x})";
 by (rtac lam_type 1);
 by (dtac apply_type 1 THEN (assume_tac 1));
@@ -180,7 +180,7 @@
 by (Asm_full_simp_tac 1);
 val lemma4 = result();
 
-Goalw AC_defs "!!Z. AC5 ==> AC4";
+Goalw AC_defs "AC5 ==> AC4";
 by (REPEAT (resolve_tac [allI,impI] 1));
 by (REPEAT (eresolve_tac [allE,ballE] 1));
 by (eresolve_tac [lemma1 RSN (2, notE)] 2 THEN (assume_tac 2));
--- a/src/ZF/AC/AC7_AC9.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/AC7_AC9.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -17,26 +17,26 @@
 by (Fast_tac 1);
 qed "mem_not_eq_not_mem";
 
-Goal "!!A. [| 0~:A; B:A |] ==> (nat->Union(A))*B ~= 0";
+Goal "[| 0~:A; B:A |] ==> (nat->Union(A))*B ~= 0";
 by (fast_tac (claset() addSDs [Sigma_empty_iff RS iffD1]
                 addDs [fun_space_emptyD, mem_not_eq_not_mem]
                 addEs [equals0D]
                 addSIs [equals0I,UnionI]) 1);
 qed "Sigma_fun_space_not0";
 
-Goal "!!A C. (ALL B:A. B eqpoll C) ==> (ALL B1:A. ALL B2:A. B1 eqpoll B2)";
+Goal "(ALL B:A. B eqpoll C) ==> (ALL B1:A. ALL B2:A. B1 eqpoll B2)";
 by (REPEAT (rtac ballI 1));
 by (resolve_tac [bspec RS (bspec RSN (2, eqpoll_sym RSN (2, eqpoll_trans)))] 1
         THEN REPEAT (assume_tac 1));
 qed "all_eqpoll_imp_pair_eqpoll";
 
-Goal "!!A. [| ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a)); b:A   \
+Goal "[| ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a)); b:A   \
 \       |] ==> P(b)=R(b)";
 by (dtac bspec 1 THEN (assume_tac 1));
 by (Asm_full_simp_tac 1);
 qed "if_eqE1";
 
-Goal "!!A. ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a))  \
+Goal "ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a))  \
 \       ==> ALL a:A. a~=b --> Q(a)=S(a)";
 by (rtac ballI 1);
 by (rtac impI 1);
@@ -44,7 +44,7 @@
 by (Asm_full_simp_tac 1);
 qed "if_eqE2";
 
-Goal "!!A. [| (lam x:A. f(x))=(lam x:A. g(x)); a:A |] ==> f(a)=g(a)";
+Goal "[| (lam x:A. f(x))=(lam x:A. g(x)); a:A |] ==> f(a)=g(a)";
 by (fast_tac (claset() addDs [subsetD]
                 addSIs [lamI]
                 addEs [equalityE, lamE]) 1);
@@ -70,7 +70,7 @@
 by (asm_full_simp_tac (simpset() addsimps [succ_not_0 RS if_not_P]) 1);
 val lemma = result();
 
-Goal "!!A. [| C:A; 0~:A |] ==> (nat->Union(A)) * C eqpoll (nat->Union(A))";
+Goal "[| C:A; 0~:A |] ==> (nat->Union(A)) * C eqpoll (nat->Union(A))";
 by (rtac eqpollI 1);
 by (fast_tac (claset() addSEs [prod_lepoll_self, not_sym RS not_emptyE,
                 subst_elem] addEs [swap]) 2);
@@ -83,7 +83,7 @@
 (* AC6 ==> AC7                                                            *)
 (* ********************************************************************** *)
 
-Goalw AC_defs "!!Z. AC6 ==> AC7";
+Goalw AC_defs "AC6 ==> AC7";
 by (Fast_tac 1);
 qed "AC6_AC7";
 
@@ -93,27 +93,27 @@
 (* the proof.                                                             *)
 (* ********************************************************************** *)
 
-Goal "!!y. y: (PROD B:A. Y*B) ==> (lam B:A. snd(y`B)): (PROD B:A. B)";
+Goal "y: (PROD B:A. Y*B) ==> (lam B:A. snd(y`B)): (PROD B:A. B)";
 by (fast_tac (claset() addSIs [lam_type, snd_type, apply_type]) 1);
 val lemma1_1 = result();
 
-Goal "!!A. y: (PROD B:{Y*C. C:A}. B)  \
+Goal "y: (PROD B:{Y*C. C:A}. B)  \
 \               ==> (lam B:A. y`(Y*B)): (PROD B:A. Y*B)";
 by (fast_tac (claset() addSIs [lam_type, apply_type]) 1);
 val lemma1_2 = result();
 
-Goal "!!A. (PROD B:{(nat->Union(A))*C. C:A}. B) ~= 0  \
+Goal "(PROD B:{(nat->Union(A))*C. C:A}. B) ~= 0  \
 \               ==> (PROD B:A. B) ~= 0";
 by (fast_tac (claset() addSIs [equals0I,lemma1_1, lemma1_2]
                 addSEs [equals0D]) 1);
 val lemma1 = result();
 
-Goal "!!A. 0 ~: A ==> 0 ~: {(nat -> Union(A)) * C. C:A}";
+Goal "0 ~: A ==> 0 ~: {(nat -> Union(A)) * C. C:A}";
 by (fast_tac (claset() addEs [RepFunE,
                 Sigma_fun_space_not0 RS not_sym RS notE]) 1);
 val lemma2 = result();
 
-Goalw AC_defs "!!Z. AC7 ==> AC6";
+Goalw AC_defs "AC7 ==> AC6";
 by (rtac allI 1);
 by (rtac impI 1);
 by (excluded_middle_tac "A=0" 1);
@@ -142,13 +142,13 @@
 by (Asm_full_simp_tac 1);
 val lemma1 = result();
 
-Goal "!!A. [| f: (PROD X:RepFun(A,p). X); D:A |]  \
+Goal "[| f: (PROD X:RepFun(A,p). X); D:A |]  \
 \               ==> (lam x:A. f`p(x))`D : p(D)";
 by (resolve_tac [beta RS ssubst] 1 THEN (assume_tac 1));
 by (fast_tac (claset() addSEs [apply_type]) 1);
 val lemma2 = result();
 
-Goalw AC_defs "!!Z. AC1 ==> AC8";
+Goalw AC_defs "AC1 ==> AC8";
 by (rtac allI 1);
 by (etac allE 1);
 by (rtac impI 1);
@@ -164,16 +164,16 @@
 (*    AC8 ==> AC1 and AC1 ==> AC9                                         *)
 (* ********************************************************************** *)
 
-Goal "!!A. ALL B1:A. ALL B2:A. B1 eqpoll B2 ==>  \
+Goal "ALL B1:A. ALL B2:A. B1 eqpoll B2 ==>  \
 \               ALL B:A*A. EX B1 B2. B=<B1,B2> & B1 eqpoll B2";
 by (Fast_tac 1);
 val lemma1 = result();
 
-Goal "!!f. f:bij(fst(<a,b>),snd(<a,b>)) ==> f:bij(a,b)";
+Goal "f:bij(fst(<a,b>),snd(<a,b>)) ==> f:bij(a,b)";
 by (Asm_full_simp_tac 1);
 val lemma2 = result();
 
-Goalw AC_defs "!!Z. AC8 ==> AC9";
+Goalw AC_defs "AC8 ==> AC9";
 by (rtac allI 1);
 by (rtac impI 1);
 by (etac allE 1);
@@ -196,7 +196,7 @@
         ((prod_commute_eqpoll RS eqpoll_imp_lepoll) RSN (2,lepoll_trans));
 
 
-Goal "!!A. [| 0~:A; A~=0 |]  \
+Goal "[| 0~:A; A~=0 |]  \
 \       ==> ALL B1: ({((nat->Union(A))*B)*nat. B:A}  \
 \               Un {cons(0,((nat->Union(A))*B)*nat). B:A}).  \
 \       ALL B2: ({((nat->Union(A))*B)*nat. B:A}  \
@@ -209,7 +209,7 @@
                         (2, prod_eqpoll_cong), Sigma_fun_space_eqpoll]) 1);
 val lemma1 = result();
 
-Goal "!!A. ALL B1:{(F*B)*N. B:A} Un {cons(0,(F*B)*N). B:A}.  \
+Goal "ALL B1:{(F*B)*N. B:A} Un {cons(0,(F*B)*N). B:A}.  \
 \       ALL B2:{(F*B)*N. B:A}  \
 \       Un {cons(0,(F*B)*N). B:A}. f`<B1,B2> : bij(B1, B2)  \
 \       ==> (lam B:A. snd(fst((f`<cons(0,(F*B)*N),(F*B)*N>)`0))) :  \
@@ -221,7 +221,7 @@
 by (fast_tac (claset() addSIs [fun_weaken_type, bij_is_fun]) 1);
 val lemma2 = result();
 
-Goalw AC_defs "!!Z. AC9 ==> AC1";
+Goalw AC_defs "AC9 ==> AC1";
 by (rtac allI 1);
 by (rtac impI 1);
 by (etac allE 1);
--- a/src/ZF/AC/AC_Equiv.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/AC_Equiv.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -38,7 +38,7 @@
 (*             lemmas concerning FOL and pure ZF theory                   *)
 (* ********************************************************************** *)
 
-Goal "!!X. (A->X)=0 ==> X=0";
+Goal "(A->X)=0 ==> X=0";
 by (fast_tac (claset() addSIs [equals0I] addEs [lam_type RSN (2, equals0D)]) 1);
 qed "fun_space_emptyD";
 
@@ -122,7 +122,7 @@
 by (fast_tac (claset() addSIs [equals0I] addEs [equals0D]) 1);
 qed "Sigma_empty_iff";
 
-Goalw [Finite_def] "!!n. n:nat ==> Finite(n)";
+Goalw [Finite_def] "n:nat ==> Finite(n)";
 by (fast_tac (claset() addSIs [eqpoll_refl]) 1);
 qed "nat_into_Finite";
 
@@ -137,7 +137,7 @@
 (* Another elimination rule for EX!                                       *)
 (* ********************************************************************** *)
 
-Goal "!!x. [| EX! x. P(x); P(x); P(y) |] ==> x=y";
+Goal "[| EX! x. P(x); P(x); P(y) |] ==> x=y";
 by (etac ex1E 1);
 by (res_inst_tac [("b","xa")] (sym RSN (2, trans)) 1);
 by (Fast_tac 1);
@@ -148,7 +148,7 @@
 (* image of a surjection                                                  *)
 (* ********************************************************************** *)
 
-Goalw [surj_def] "!!f. f : surj(A, B) ==> f``A = B";
+Goalw [surj_def] "f : surj(A, B) ==> f``A = B";
 by (etac CollectE 1);
 by (resolve_tac [subset_refl RSN (2, image_fun) RS ssubst] 1 
     THEN (assume_tac 1));
@@ -156,11 +156,11 @@
 qed "surj_image_eq";
 
 
-Goal "!!y. succ(x) lepoll y ==> y ~= 0";
+Goal "succ(x) lepoll y ==> y ~= 0";
 by (fast_tac (claset() addSDs [lepoll_0_is_0]) 1);
 qed "succ_lepoll_imp_not_empty";
 
-Goal "!!x. x eqpoll succ(n) ==> x ~= 0";
+Goal "x eqpoll succ(n) ==> x ~= 0";
 by (fast_tac (claset() addSEs [eqpoll_sym RS eqpoll_0_is_0 RS succ_neq_0]) 1);
 qed "eqpoll_succ_imp_not_empty";
 
--- a/src/ZF/AC/Cardinal_aux.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/Cardinal_aux.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -26,7 +26,7 @@
 by (fast_tac (claset() addSDs [lepoll_imp_ex_le_eqpoll] addSEs [leE]) 1);
 qed "lesspoll_imp_ex_lt_eqpoll";
 
-Goalw [InfCard_def] "!!i. [| ~Finite(i); Ord(i) |] ==> InfCard(|i|)";
+Goalw [InfCard_def] "[| ~Finite(i); Ord(i) |] ==> InfCard(|i|)";
 by (rtac conjI 1);
 by (rtac Card_cardinal 1);
 by (resolve_tac [Card_nat RS (Card_def RS def_imp_iff RS iffD1 RS ssubst)] 1);
@@ -39,12 +39,12 @@
         asm_rl, lepoll_refl] MRS (prod_lepoll_mono RSN (2, lepoll_trans))
         |> standard;
 
-Goal "!!A. [| A lepoll B; 2 lepoll A |] ==> A+B lepoll A*B";
+Goal "[| A lepoll B; 2 lepoll A |] ==> A+B lepoll A*B";
 by (REPEAT (ares_tac [[sum_lepoll_mono, sum_lepoll_prod]
                 MRS lepoll_trans, lepoll_refl] 1));
 qed "lepoll_imp_sum_lepoll_prod";
 
-Goal "!!A. [| A eqpoll i; B eqpoll i; ~Finite(i); Ord(i) |]  \
+Goal "[| A eqpoll i; B eqpoll i; ~Finite(i); Ord(i) |]  \
 \               ==> A Un B eqpoll i";
 by (rtac eqpollI 1);
 by (eresolve_tac [subset_imp_lepoll RSN (2, eqpoll_sym RS eqpoll_imp_lepoll
@@ -73,7 +73,6 @@
 qed "double_Diff_sing";
 
 goal ZF.thy "if({y,z}-{z}=0, z, THE w. {y,z}-{z}={w}) = y";
-by (split_tac [split_if] 1);
 by (asm_full_simp_tac (simpset() addsimps [double_Diff_sing, Diff_eq_0_iff]) 1);
 by (fast_tac (claset() addSIs [the_equality] addEs [equalityE]) 1);
 qed "paired_bij_lemma";
@@ -81,19 +80,20 @@
 Goal "(lam y:{{y,z}. y:x}. if(y-{z}=0, z, THE w. y-{z}={w}))  \
 \               : bij({{y,z}. y:x}, x)";
 by (res_inst_tac [("d","%a. {a,z}")] lam_bijective 1);
-by (TRYALL (fast_tac (claset() addSEs [RepFunE] addSIs [RepFunI] 
-                addss (simpset() addsimps [paired_bij_lemma]))));
+by (auto_tac (claset(),
+	      simpset() delsplits [split_if]
+		        addsimps [paired_bij_lemma]));
 qed "paired_bij";
 
 Goalw [eqpoll_def] "{{y,z}. y:x} eqpoll x";
 by (fast_tac (claset() addSIs [paired_bij]) 1);
 qed "paired_eqpoll";
 
-Goal "!!A. EX B. B eqpoll A & B Int C = 0";
+Goal "EX B. B eqpoll A & B Int C = 0";
 by (fast_tac (claset() addSIs [paired_eqpoll, equals0I] addEs [mem_asym]) 1);
 qed "ex_eqpoll_disjoint";
 
-Goal "!!A. [| A lepoll i; B lepoll i; ~Finite(i); Ord(i) |]  \
+Goal "[| A lepoll i; B lepoll i; ~Finite(i); Ord(i) |]  \
 \               ==> A Un B lepoll i";
 by (res_inst_tac [("A1","i"), ("C1","i")] (ex_eqpoll_disjoint RS exE) 1);
 by (etac conjE 1);
@@ -105,7 +105,7 @@
         THEN (REPEAT (assume_tac 1)));
 qed "Un_lepoll_Inf_Ord";
 
-Goal "!!P. [| P(i); i:j; Ord(j) |] ==> (LEAST i. P(i)) : j";
+Goal "[| P(i); i:j; Ord(j) |] ==> (LEAST i. P(i)) : j";
 by (eresolve_tac [Least_le RS leE] 1);
 by (etac Ord_in_Ord 1 THEN (assume_tac 1));
 by (etac ltE 1);
@@ -113,7 +113,7 @@
 by (etac subst_elem 1 THEN (assume_tac 1));
 qed "Least_in_Ord";
 
-Goal "!!x. [| well_ord(x,r); y<=x; y lepoll succ(n); n:nat |]  \
+Goal "[| well_ord(x,r); y<=x; y lepoll succ(n); n:nat |]  \
 \       ==> y-{THE b. first(b,y,r)} lepoll n";
 by (res_inst_tac [("Q","y=0")] (excluded_middle RS disjE) 1);
 by (fast_tac (claset() addSIs [Diff_sing_lepoll, the_first_in]) 1);
@@ -126,14 +126,14 @@
 by (Fast_tac 1);
 qed "UN_subset_split";
 
-Goalw [lepoll_def] "!!a. Ord(a) ==> (UN x:a. {P(x)}) lepoll a";
+Goalw [lepoll_def] "Ord(a) ==> (UN x:a. {P(x)}) lepoll a";
 by (res_inst_tac [("x","lam z:(UN x:a. {P(x)}). (LEAST i. P(i)=z)")] exI 1);
 by (res_inst_tac [("d","%z. P(z)")] lam_injective 1);
 by (fast_tac (claset() addSIs [Least_in_Ord]) 1);
 by (fast_tac (claset() addIs [LeastI] addSEs [Ord_in_Ord]) 1);
 qed "UN_sing_lepoll";
 
-Goal "!!a T. [| well_ord(T, R); ~Finite(a); Ord(a); n:nat |] ==>  \
+Goal "[| well_ord(T, R); ~Finite(a); Ord(a); n:nat |] ==>  \
 \       ALL f. (ALL b:a. f`b lepoll n & f`b <= T) --> (UN b:a. f`b) lepoll a";
 by (etac nat_induct 1);
 by (rtac allI 1);
@@ -155,13 +155,13 @@
 by (etac UN_sing_lepoll 1);
 qed "UN_fun_lepoll_lemma";
 
-Goal "!!a f. [| ALL b:a. f`b lepoll n & f`b <= T; well_ord(T, R);  \
+Goal "[| ALL b:a. f`b lepoll n & f`b <= T; well_ord(T, R);  \
 \       ~Finite(a); Ord(a); n:nat |] ==> (UN b:a. f`b) lepoll a";
 by (eresolve_tac [UN_fun_lepoll_lemma RS allE] 1 THEN (REPEAT (assume_tac 1)));
 by (Fast_tac 1);
 qed "UN_fun_lepoll";
 
-Goal "!!a f. [| ALL b:a. F(b) lepoll n & F(b) <= T; well_ord(T, R);  \
+Goal "[| ALL b:a. F(b) lepoll n & F(b) <= T; well_ord(T, R);  \
 \       ~Finite(a); Ord(a); n:nat |] ==> (UN b:a. F(b)) lepoll a";
 by (rtac impE 1 THEN (assume_tac 3));
 by (res_inst_tac [("f","lam b:a. F(b)")] (UN_fun_lepoll) 2 
@@ -170,7 +170,7 @@
 by (Asm_full_simp_tac 1);
 qed "UN_lepoll";
 
-Goal "!!a. Ord(a) ==> (UN b:a. F(b)) = (UN b:a. F(b) - (UN c:b. F(c)))";
+Goal "Ord(a) ==> (UN b:a. F(b)) = (UN b:a. F(b) - (UN c:b. F(c)))";
 by (rtac equalityI 1);
 by (Fast_tac 2);
 by (rtac subsetI 1);
@@ -185,7 +185,7 @@
 by (eresolve_tac [Ord_Least RSN (2, ltI)] 1);
 qed "UN_eq_UN_Diffs";
 
-Goalw [eqpoll_def] "!!A B. A Int B = 0 ==> A Un B eqpoll A + B";
+Goalw [eqpoll_def] "A Int B = 0 ==> A Un B eqpoll A + B";
 by (res_inst_tac [("x","lam a:A Un B. if(a:A,Inl(a),Inr(a))")] exI 1);
 by (res_inst_tac [("d","%z. case(%x. x, %x. x, z)")] lam_bijective 1);
 by (fast_tac (claset() addSIs [if_type, InlI, InrI]) 1);
@@ -207,7 +207,7 @@
 (* Diff_lesspoll_eqpoll_Card                                              *)
 (* ********************************************************************** *)
 
-Goal "!!A B. [| A eqpoll a; ~Finite(a); Card(a); B lesspoll a;  \
+Goal "[| A eqpoll a; ~Finite(a); Card(a); B lesspoll a;  \
 \               A-B lesspoll a |] ==> P";
 by (REPEAT (eresolve_tac [lesspoll_imp_ex_lt_eqpoll RS exE,
         Card_is_Ord, conjE] 1));
@@ -237,7 +237,7 @@
 by (fast_tac (claset() addSDs [lesspoll_def RS def_imp_iff RS iffD1]) 1);
 qed "Diff_lesspoll_eqpoll_Card_lemma";
 
-Goal "!!A. [| A eqpoll a; ~Finite(a); Card(a); B lesspoll a |]  \
+Goal "[| A eqpoll a; ~Finite(a); Card(a); B lesspoll a |]  \
 \       ==> A - B eqpoll a";
 by (rtac swap 1 THEN (Fast_tac 1));
 by (rtac Diff_lesspoll_eqpoll_Card_lemma 1 THEN (REPEAT (assume_tac 1)));
--- a/src/ZF/AC/DC.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/DC.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -38,7 +38,7 @@
 by (Fast_tac 1);
 val lemma1_1 = result();
 
-Goal "!!X. ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R)  \
+Goal "ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R)  \
 \       ==> {<z1,z2>: (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) *  \
 \               (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}).  \
 \               domain(z2)=succ(domain(z1))  \
@@ -58,7 +58,7 @@
                 singleton_0]) 1);
 val lemma1_2 = result();
 
-Goal "!!X. ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R)  \
+Goal "ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R)  \
 \       ==> range({<z1,z2>: (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) *  \
 \               (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}).  \
 \               domain(z2)=succ(domain(z1))  \
@@ -92,7 +92,7 @@
         addsimps [domain_cons, domain_of_fun, succ_def, restrict_cons_eq]) 1);
 val lemma1_3 = result();
 
-Goal "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R});  \
+Goal "[| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R});  \
 \       RR = {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1))  \
 \       & restrict(z2, domain(z1)) = z1};  \
 \       ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R)  \
@@ -162,12 +162,12 @@
         addSEs [bexE, nat_into_Ord RSN (2, OrdmemD) RS subsetD]) 1);
 val lemma3_1 = result();
 
-Goal "!!n. ALL x:n. f`succ(n)`x = f`succ(x)`x   \
+Goal "ALL x:n. f`succ(n)`x = f`succ(x)`x   \
 \       ==> {f`succ(x)`x. x:n} = {f`succ(n)`x. x:n}";
 by (Asm_full_simp_tac 1);
 val lemma3_2 = result();
 
-Goal "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R});  \
+Goal "[| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R});  \
 \       ALL n:nat. <f`n, f`succ(n)> :  \
 \       {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1))  \
 \       & restrict(z2, domain(z1)) = z1};  \
@@ -186,7 +186,7 @@
                             (2, image_fun RS sym)]) 1);
 val lemma3 = result();
 
-Goal "!!f. [| f:A->B; B<=C |] ==> f:A->C";
+Goal "[| f:A->B; B<=C |] ==> f:A->C";
 by (rtac Pi_type 1 THEN (assume_tac 1));
 by (fast_tac (claset() addSEs [apply_type]) 1);
 qed "fun_type_gen";
@@ -247,7 +247,7 @@
         addSIs [Ord_nat]) 1);
 qed "lesspoll_nat_is_Finite";
 
-Goal "!!n. n:nat ==> ALL A. (A eqpoll n & A <= X) --> A : Fin(X)";
+Goal "n:nat ==> ALL A. (A eqpoll n & A <= X) --> A : Fin(X)";
 by (etac nat_induct 1);
 by (rtac allI 1);
 by (fast_tac (claset() addSIs [Fin.emptyI]
@@ -266,7 +266,7 @@
 by (asm_full_simp_tac (simpset() addsimps [cons_Diff]) 1);
 qed "Finite_Fin_lemma";
 
-Goalw [Finite_def] "!!A. [| Finite(A); A <= X |] ==> A : Fin(X)";
+Goalw [Finite_def] "[| Finite(A); A <= X |] ==> A : Fin(X)";
 by (etac bexE 1);
 by (dtac Finite_Fin_lemma 1);
 by (etac allE 1);
@@ -275,7 +275,7 @@
 by (Fast_tac 1);
 qed "Finite_Fin";
 
-Goal "!!x. x: X  \
+Goal "x: X  \
 \ ==> {<0,x>}: (UN n:nat. {f:succ(n)->X. ALL k:n. <f`k, f`succ(k)> : R})";
 by (rtac (nat_0I RS UN_I) 1);
 by (fast_tac (claset() addSIs [singleton_fun RS Pi_type]
@@ -307,31 +307,31 @@
                                      cons_fun_type2, empty_fun]) 1);
 val lemma4 = result();
 
-Goal "!!f. [| f:nat->X; n:nat |] ==>  \
+Goal "[| f:nat->X; n:nat |] ==>  \
 \       (UN x:f``succ(n). P(x)) =  P(f`n) Un (UN x:f``n. P(x))";
 by (asm_full_simp_tac (simpset()
         addsimps [Ord_nat RSN (2, OrdmemD) RSN (2, image_fun),
         [nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1);
 qed "UN_image_succ_eq";
 
-Goal "!!f. [| (UN x:f``n. P(x)) = y; P(f`n) = succ(y);  \
+Goal "[| (UN x:f``n. P(x)) = y; P(f`n) = succ(y);  \
 \       f:nat -> X; n:nat |] ==> (UN x:f``succ(n). P(x)) = succ(y)";
 by (asm_full_simp_tac (simpset() addsimps [UN_image_succ_eq]) 1);
 by (Fast_tac 1);
 qed "UN_image_succ_eq_succ";
 
-Goal "!!f. [| f:succ(n) -> D;  n:nat;  \
+Goal "[| f:succ(n) -> D;  n:nat;  \
 \       domain(f)=succ(x); x=y |] ==> f`y : D";
 by (fast_tac (claset() addEs [apply_type]
         addSDs [[sym, domain_of_fun] MRS trans]) 1);
 qed "apply_domain_type";
 
-Goal "!!f. [| f : nat -> X; n:nat |] ==> f``succ(n) = cons(f`n, f``n)";
+Goal "[| f : nat -> X; n:nat |] ==> f``succ(n) = cons(f`n, f``n)";
 by (asm_full_simp_tac (simpset()
         addsimps [nat_succI, Ord_nat RSN (2, OrdmemD), image_fun]) 1);
 qed "image_fun_succ";
 
-Goal "!!f. [| domain(f`n) = succ(u); f : nat -> (UN n:nat.  \
+Goal "[| domain(f`n) = succ(u); f : nat -> (UN n:nat.  \
 \       {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R});  \
 \       u=k; n:nat  \
 \       |] ==> f`n : succ(k) -> domain(R)";
@@ -339,7 +339,7 @@
 by (fast_tac (claset() addEs [UN_E, domain_eq_imp_fun_type]) 1);
 qed "f_n_type";
 
-Goal "!!f. [| f : nat -> (UN n:nat.  \
+Goal "[| f : nat -> (UN n:nat.  \
 \       {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R});  \
 \       domain(f`n) = succ(k); n:nat  \
 \       |] ==> ALL i:k. <f`n`i, f`n`succ(i)> : R";
@@ -360,7 +360,7 @@
 by (asm_full_simp_tac (simpset() addsimps [subsetD RS cons_val_k]) 1);
 qed "restrict_cons_eq_restrict";
 
-Goal "!!f. [| ALL x:f``n. restrict(f`n, domain(x))=x;  \
+Goal "[| ALL x:f``n. restrict(f`n, domain(x))=x;  \
 \       f : nat -> (UN n:nat.  \
 \       {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R});  \
 \       n:nat; domain(f`n) = succ(n);  \
@@ -436,7 +436,7 @@
 qed "simplify_recursion";
 
 
-Goal "!!X. [| XX = (UN n:nat.  \
+Goal "[| XX = (UN n:nat.  \
 \               {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R});  \
 \       ALL b<nat. <f``b, f`b> :  \
 \       {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f))  \
@@ -518,7 +518,7 @@
 by (REPEAT (fast_tac (claset() addDs [not_eq, not_eq RS not_sym]) 1));
 qed "fun_Ord_inj";
 
-Goal "!!a. [| f:X->Y; A<=X; a:A |] ==> f`a : f``A";
+Goal "[| f:X->Y; A<=X; a:A |] ==> f`a : f``A";
 by (fast_tac (claset() addSEs [image_fun RS ssubst]) 1);
 qed "value_in_image";
 
@@ -562,7 +562,7 @@
 by (Asm_full_simp_tac 1);
 qed "lam_images_eq";
 
-Goalw [lesspoll_def] "!!K. [| Card(K); b:K |] ==> b lesspoll K";
+Goalw [lesspoll_def] "[| Card(K); b:K |] ==> b lesspoll K";
 by (asm_full_simp_tac (simpset() addsimps [Card_iff_initial]) 1);
 by (fast_tac (claset() addSIs [le_imp_lepoll, ltI, leI]) 1);
 qed "in_Card_imp_lesspoll";
@@ -571,14 +571,14 @@
 by (fast_tac (claset() addSIs [lam_type, RepFunI]) 1);
 qed "lam_type_RepFun";
 
-Goal "!!Z. [| ALL Y:Pow(X). Y lesspoll a --> (EX x:X. <Y, x> : R);  \
+Goal "[| ALL Y:Pow(X). Y lesspoll a --> (EX x:X. <Y, x> : R);  \
 \       b:a; Z:Pow(X); Z lesspoll a |]  \
 \       ==> {x:X. <Z,x> : R} ~= 0";
 by (fast_tac (FOL_cs addEs [bexE, equals0D]
         addSDs [bspec] addIs [CollectI]) 1);
 val lemma_ = result();
 
-Goal "!!K. [| Card(K); well_ord(X,Q);  \
+Goal "[| Card(K); well_ord(X,Q);  \
 \       ALL Y:Pow(X). Y lesspoll K --> (EX x:X. <Y, x> : R); b:K |]  \
 \       ==> ff(b, X, Q, R) : {x:X. <(lam c:b. ff(c, X, Q, R))``b, x> : R}";
 by (res_inst_tac [("P","b:K")] impE 1 THEN TRYALL assume_tac);
--- a/src/ZF/AC/DC_lemmas.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/DC_lemmas.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -14,7 +14,7 @@
 by (REPEAT (eresolve_tac [RepFunE, LeastI, prem RS Ord_in_Ord] 1));
 qed "RepFun_lepoll";
 
-Goalw [lesspoll_def] "!!n. n:nat ==> n lesspoll nat";
+Goalw [lesspoll_def] "n:nat ==> n lesspoll nat";
 by (rtac conjI 1);
 by (eresolve_tac [Ord_nat RSN (2, OrdmemD) RS subset_imp_lepoll] 1);
 by (rtac notI 1);
@@ -46,35 +46,35 @@
 by (fast_tac (claset() addSEs [mem_irrefl]) 1);
 qed "mem_not_eq";
 
-Goalw [succ_def] "!!g. g:n->X ==> cons(<n,x>, g) : succ(n) -> cons(x, X)";
+Goalw [succ_def] "g:n->X ==> cons(<n,x>, g) : succ(n) -> cons(x, X)";
 by (fast_tac (claset() addSIs [fun_extend] addSEs [mem_irrefl]) 1);
 qed "cons_fun_type";
 
-Goal "!!g. [| g:n->X; x:X |] ==> cons(<n,x>, g) : succ(n) -> X";
+Goal "[| g:n->X; x:X |] ==> cons(<n,x>, g) : succ(n) -> X";
 by (etac (cons_absorb RS subst) 1 THEN etac cons_fun_type 1);
 qed "cons_fun_type2";
 
-Goal "!!n. n: nat ==> cons(<n,x>, g)``n = g``n";
+Goal "n: nat ==> cons(<n,x>, g)``n = g``n";
 by (fast_tac (claset() addSEs [mem_irrefl]) 1);
 qed "cons_image_n";
 
-Goal "!!n. g:n->X ==> cons(<n,x>, g)`n = x";
+Goal "g:n->X ==> cons(<n,x>, g)`n = x";
 by (fast_tac (claset() addSIs [apply_equality] addSEs [cons_fun_type]) 1);
 qed "cons_val_n";
 
-Goal "!!k. k : n ==> cons(<n,x>, g)``k = g``k";
+Goal "k : n ==> cons(<n,x>, g)``k = g``k";
 by (fast_tac (claset() addEs [mem_asym]) 1);
 qed "cons_image_k";
 
-Goal "!!k. [| k:n; g:n->X |] ==> cons(<n,x>, g)`k = g`k";
+Goal "[| k:n; g:n->X |] ==> cons(<n,x>, g)`k = g`k";
 by (fast_tac (claset() addSIs [apply_equality, consI2] addSEs [cons_fun_type, apply_Pair]) 1);
 qed "cons_val_k";
 
-Goal "!!f. domain(f)=x ==> domain(cons(<x,y>, f)) = succ(x)";
+Goal "domain(f)=x ==> domain(cons(<x,y>, f)) = succ(x)";
 by (asm_full_simp_tac (simpset() addsimps [domain_cons, succ_def]) 1);
 qed "domain_cons_eq_succ";
 
-Goalw [restrict_def] "!!g. g:n->X ==> restrict(cons(<n,x>, g), n)=g";
+Goalw [restrict_def] "g:n->X ==> restrict(cons(<n,x>, g), n)=g";
 by (rtac fun_extension 1);
 by (rtac lam_type 1);
 by (eresolve_tac [cons_fun_type RS apply_type] 1);
@@ -83,7 +83,7 @@
 by (asm_full_simp_tac (simpset() addsimps [cons_val_k]) 1);
 qed "restrict_cons_eq";
 
-Goal "!!k. [| Ord(k); i:k |] ==> succ(i) : succ(k)";
+Goal "[| Ord(k); i:k |] ==> succ(i) : succ(k)";
 by (resolve_tac [Ord_linear RS disjE] 1 THEN (assume_tac 3));
 by (REPEAT (fast_tac (claset() addSIs [Ord_succ]
         addEs [Ord_in_Ord, mem_irrefl, mem_asym]
@@ -96,12 +96,12 @@
 by (Asm_full_simp_tac 1);
 qed "restrict_eq_imp_val_eq";
 
-Goal "!!f. [| domain(f)=A; f:B->C |] ==> f:A->C";
+Goal "[| domain(f)=A; f:B->C |] ==> f:A->C";
 by (forward_tac [domain_of_fun] 1);
 by (Fast_tac 1);
 qed "domain_eq_imp_fun_type";
 
-Goal "!!R. [| R <= A * B; R ~= 0 |] ==> EX x. x:domain(R)";
+Goal "[| R <= A * B; R ~= 0 |] ==> EX x. x:domain(R)";
 by (fast_tac (claset() addSEs [not_emptyE]) 1);
 qed "ex_in_domain";
 
--- a/src/ZF/AC/HH.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/HH.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -23,16 +23,15 @@
 
 Goal "HH(f,x,a) : Pow(x)-{0} | HH(f,x,a)={x}";
 by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
-by (simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI] 
-                    addsplits [split_if]) 1);
+by (simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]) 1);
 by (Fast_tac 1);
 qed "HH_values";
 
-Goal "!!A. B<=A ==> X-(UN a:A. P(a)) = X-(UN a:A-B. P(a))-(UN b:B. P(b))";
+Goal "B<=A ==> X-(UN a:A. P(a)) = X-(UN a:A-B. P(a))-(UN b:B. P(b))";
 by (Fast_tac 1);
 qed "subset_imp_Diff_eq";
 
-Goal "!!c. [| c:a-b; b<a |] ==> c=b | b<c & c<a";
+Goal "[| c:a-b; b<a |] ==> c=b | b<c & c<a";
 by (etac ltE 1);
 by (dtac Ord_linear 1);
 by (fast_tac (claset() addSIs [ltI] addIs [Ord_in_Ord]) 2);
@@ -44,14 +43,14 @@
 by (fast_tac (claset() addSDs [prem] addSEs [mem_irrefl]) 1);
 qed "Diff_UN_eq_self";
 
-Goal "!!a. x - (UN b:a. HH(f,x,b)) = x - (UN b:a1. HH(f,x,b))  \
+Goal "x - (UN b:a. HH(f,x,b)) = x - (UN b:a1. HH(f,x,b))  \
 \               ==> HH(f,x,a) = HH(f,x,a1)";
 by (resolve_tac [HH_def_satisfies_eq RS
                 (HH_def_satisfies_eq RS sym RSN (3, trans RS trans))] 1);
 by (etac subst_context 1);
 qed "HH_eq";
 
-Goal "!!a. [| HH(f,x,b)={x}; b<a |] ==> HH(f,x,a)={x}";
+Goal "[| HH(f,x,b)={x}; b<a |] ==> HH(f,x,a)={x}";
 by (res_inst_tac [("P","b<a")] impE 1 THEN REPEAT (assume_tac 2));
 by (eresolve_tac [lt_Ord2 RS trans_induct] 1);
 by (rtac impI 1);
@@ -64,31 +63,31 @@
 by (fast_tac (claset() addEs [ltE]) 1);
 qed "HH_is_x_gt_too";
 
-Goal "!!a. [| HH(f,x,a) : Pow(x)-{0}; b<a |] ==> HH(f,x,b) : Pow(x)-{0}";
+Goal "[| HH(f,x,a) : Pow(x)-{0}; b<a |] ==> HH(f,x,b) : Pow(x)-{0}";
 by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 1));
 by (dtac HH_is_x_gt_too 1 THEN (assume_tac 1));
 by (dtac subst 1 THEN (assume_tac 1));
 by (fast_tac (claset() addSEs [mem_irrefl]) 1);
 qed "HH_subset_x_lt_too";
 
-Goal "!!a. HH(f,x,a) : Pow(x)-{0}   \
+Goal "HH(f,x,a) : Pow(x)-{0}   \
 \               ==> HH(f,x,a) : Pow(x - (UN b:a. HH(f,x,b)))-{0}";
 by (dresolve_tac [HH_def_satisfies_eq RS subst] 1);
 by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
 by (asm_full_simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]) 1);
 by (dresolve_tac [split_if RS iffD1] 1);
-by (simp_tac (simpset() addsplits [split_if] ) 1);
+by (Simp_tac 1);
 by (fast_tac (subset_cs addSEs [mem_irrefl]) 1);
 qed "HH_subset_x_imp_subset_Diff_UN";
 
-Goal "!!x. [| HH(f,x,v)=HH(f,x,w); HH(f,x,v): Pow(x)-{0}; v:w |] ==> P";
+Goal "[| HH(f,x,v)=HH(f,x,w); HH(f,x,v): Pow(x)-{0}; v:w |] ==> P";
 by (forw_inst_tac [("P","%y. y: Pow(x)-{0}")] subst 1 THEN (assume_tac 1));
 by (dres_inst_tac [("a","w")] HH_subset_x_imp_subset_Diff_UN 1);
 by (dtac subst_elem 1 THEN (assume_tac 1));
 by (fast_tac (claset() addSIs [singleton_iff RS iffD2, equals0I]) 1);
 qed "HH_eq_arg_lt";
 
-Goal "!!x. [| HH(f,x,v)=HH(f,x,w); HH(f,x,w): Pow(x)-{0};  \
+Goal "[| HH(f,x,v)=HH(f,x,w); HH(f,x,w): Pow(x)-{0};  \
 \               Ord(v); Ord(w) |] ==> v=w";
 by (res_inst_tac [("j","w")] Ord_linear_lt 1 THEN TRYALL assume_tac);
 by (resolve_tac [sym RS (ltD RSN (3, HH_eq_arg_lt))] 2
@@ -115,7 +114,7 @@
 by (fast_tac (claset() addSIs [Ord_Hartog, HH_Hartog_is_x, LeastI]) 1);
 qed "HH_Least_eq_x";
 
-Goal "!!a. a:(LEAST i. HH(f,x,i)={x}) ==> HH(f,x,a) : Pow(x)-{0}";
+Goal "a:(LEAST i. HH(f,x,i)={x}) ==> HH(f,x,a) : Pow(x)-{0}";
 by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 1));
 by (rtac less_LeastE 1);
 by (eresolve_tac [Ord_Least RSN (2, ltI)] 2);
@@ -134,7 +133,7 @@
                 addSEs [HH_eq_imp_arg_eq, Ord_Least RS Ord_in_Ord]) 1);
 qed "lam_Least_HH_inj_Pow";
 
-Goal "!!x. ALL a:(LEAST i. HH(f,x,i)={x}). EX z:x. HH(f,x,a) = {z}  \
+Goal "ALL a:(LEAST i. HH(f,x,i)={x}). EX z:x. HH(f,x,a) = {z}  \
 \               ==> (lam a:(LEAST i. HH(f,x,i)={x}). HH(f,x,a))  \
 \                       : inj(LEAST i. HH(f,x,i)={x}, {{y}. y:x})";
 by (resolve_tac [lam_Least_HH_inj_Pow RS inj_strengthen_type] 1);
@@ -151,12 +150,12 @@
 by (deepen_tac (claset() addSIs [bexI] addSEs [equalityE]) 4 1);
 qed "lam_surj_sing";
 
-Goal "!!x. y:Pow(x)-{0} ==> x ~= 0";
+Goal "y:Pow(x)-{0} ==> x ~= 0";
 by (fast_tac (claset() addSIs [equals0I, singletonI RS subst_elem]
                 addSDs [equals0D]) 1);
 qed "not_emptyI2";
 
-Goal "!!f. f`(x - (UN j:i. HH(f,x,j))): Pow(x - (UN j:i. HH(f,x,j)))-{0}  \
+Goal "f`(x - (UN j:i. HH(f,x,j))): Pow(x - (UN j:i. HH(f,x,j)))-{0}  \
 \       ==> HH(f, x, i) : Pow(x) - {0}";
 by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
 by (asm_full_simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI,
@@ -176,8 +175,7 @@
 
 Goal "HH(f,x,i)=f`(x - (UN j:i. HH(f,x,j))) | HH(f,x,i)={x}";
 by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
-by (simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]
-              addsplits [split_if]) 1);
+by (simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]) 1);
 qed "HH_values2";
 
 Goal
@@ -187,7 +185,7 @@
         addSDs [singleton_subsetD]) 1);
 qed "HH_subset_imp_eq";
 
-Goal "!!f. [| f : (Pow(x)-{0}) -> {{z}. z:x};  \
+Goal "[| f : (Pow(x)-{0}) -> {{z}. z:x};  \
 \       a:(LEAST i. HH(f,x,i)={x}) |] ==> EX z:x. HH(f,x,a) = {z}";
 by (dtac less_Least_subset_x 1);
 by (forward_tac [HH_subset_imp_eq] 1);
@@ -207,7 +205,7 @@
                               f_sing_imp_HH_sing]) 1);
 qed "f_sing_lam_bij";
 
-Goal "!!f. f: (PROD X: Pow(x)-{0}. F(X))  \
+Goal "f: (PROD X: Pow(x)-{0}. F(X))  \
 \       ==> (lam X:Pow(x)-{0}. {f`X}) : (PROD X: Pow(x)-{0}. {{z}. z:F(X)})";
 by (fast_tac (FOL_cs addSIs [lam_type, RepFun_eqI, singleton_eq_iff RS iffD2]
                      addDs [apply_type]) 1);
--- a/src/ZF/AC/Hartog.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/Hartog.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -7,12 +7,12 @@
 
 open Hartog;
 
-Goal "!!X. ALL a. Ord(a) --> a:X ==> P";
+Goal "ALL a. Ord(a) --> a:X ==> P";
 by (res_inst_tac [("X1","{y:X. Ord(y)}")] (ON_class RS revcut_rl) 1);
 by (Fast_tac 1);
 qed "Ords_in_set";
 
-Goalw [lepoll_def] "!!X. [| Ord(a); a lepoll X |] ==>  \
+Goalw [lepoll_def] "[| Ord(a); a lepoll X |] ==>  \
 \               EX Y. Y<=X & (EX R. well_ord(Y,R) & ordertype(Y,R)=a)";
 by (etac exE 1);
 by (REPEAT (resolve_tac [exI, conjI] 1));
@@ -28,7 +28,7 @@
         THEN (REPEAT (assume_tac 1)));
 qed "Ord_lepoll_imp_ex_well_ord";
 
-Goal "!!X. [| Ord(a); a lepoll X |] ==>  \
+Goal "[| Ord(a); a lepoll X |] ==>  \
 \               EX Y. Y<=X & (EX R. R<=X*X & ordertype(Y,R)=a)";
 by (dtac Ord_lepoll_imp_ex_well_ord 1 THEN (assume_tac 1));
 by Safe_tac;
@@ -37,7 +37,7 @@
 by (Fast_tac 1);
 qed "Ord_lepoll_imp_eq_ordertype";
 
-Goal "!!X. ALL a. Ord(a) --> a lepoll X ==>  \
+Goal "ALL a. Ord(a) --> a lepoll X ==>  \
 \       ALL a. Ord(a) -->  \
 \       a:{a. Z:Pow(X)*Pow(X*X), EX Y R. Z=<Y,R> & ordertype(Y,R)=a}";
 by (REPEAT (resolve_tac [allI,impI] 1));
@@ -47,7 +47,7 @@
 by (fast_tac (claset() addSIs [ReplaceI] addEs [sym]) 1);
 qed "Ords_lepoll_set_lemma";
 
-Goal "!!X. ALL a. Ord(a) --> a lepoll X ==> P";
+Goal "ALL a. Ord(a) --> a lepoll X ==> P";
 by (eresolve_tac [Ords_lepoll_set_lemma RS Ords_in_set] 1);
 qed "Ords_lepoll_set";
 
@@ -70,11 +70,11 @@
 by (rtac Ord_Least 1);
 qed "Ord_Hartog";
 
-Goalw [Hartog_def] "!!i. [| i < Hartog(A); ~ i lepoll A |] ==> P";
+Goalw [Hartog_def] "[| i < Hartog(A); ~ i lepoll A |] ==> P";
 by (fast_tac (claset() addEs [less_LeastE]) 1);
 qed "less_HartogE1";
 
-Goal "!!i. [| i < Hartog(A); i eqpoll Hartog(A) |] ==> P";
+Goal "[| i < Hartog(A); i eqpoll Hartog(A) |] ==> P";
 by (fast_tac (claset() addEs [less_HartogE1, eqpoll_sym RS eqpoll_imp_lepoll
                 RS lepoll_trans RS HartogE]) 1);
 qed "less_HartogE";
--- a/src/ZF/AC/WO1_AC.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/WO1_AC.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -39,7 +39,7 @@
 (* WO1 ==> AC10(n) (n >= 1)                                               *)
 (* ********************************************************************** *)
 
-Goalw [WO1_def] "!!A. [| WO1; ALL B:A. EX C:D(B). P(C,B) |]  \
+Goalw [WO1_def] "[| WO1; ALL B:A. EX C:D(B). P(C,B) |]  \
 \               ==> EX f. ALL B:A. P(f`B,B)";
 by (eres_inst_tac [("x","Union({{C:D(B). P(C,B)}. B:A})")] allE 1);
 by (etac exE 1);
@@ -52,7 +52,7 @@
                 addSEs [CollectD2]) 1);
 val lemma1 = result();
 
-Goalw [WO1_def] "!!A. [| ~Finite(B); WO1 |] ==> |B| + |B| eqpoll  B";
+Goalw [WO1_def] "[| ~Finite(B); WO1 |] ==> |B| + |B| eqpoll  B";
 by (rtac eqpoll_trans 1);
 by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll]) 2);
 by (resolve_tac [eqpoll_sym RS eqpoll_trans] 1);
@@ -68,12 +68,12 @@
 by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll]) 1);
 val lemma2_1 = result();
 
-Goal "!!f. f : bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i:D} : Pow(Pow(B))";
+Goal "f : bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i:D} : Pow(Pow(B))";
 by (fast_tac (claset() addSIs [InlI, InrI]
                 addSEs [RepFunE, bij_is_fun RS apply_type]) 1);
 val lemma2_2 = result();
 
-Goal "!!f. [| f:inj(A,B); f`a = f`b; a:A; b:A |] ==> a=b";
+Goal "[| f:inj(A,B); f`a = f`b; a:A; b:A |] ==> a=b";
 by (rtac inj_equality 1);
 by (TRYALL (fast_tac (claset() addSEs [inj_is_fun RS apply_Pair] addEs [subst])));
 val lemma = result();
@@ -104,7 +104,7 @@
 by (fast_tac (claset() addSEs [inj_is_fun RS apply_type]) 1);
 val lemma2_5 = result();
 
-Goal "!!A. [| WO1; ~Finite(B); 1 le n  |]  \
+Goal "[| WO1; ~Finite(B); 1 le n  |]  \
 \       ==> EX C:Pow(Pow(B)). pairwise_disjoint(C) &  \
 \               sets_of_size_between(C, 2, succ(n)) &  \
 \               Union(C)=B";
@@ -114,6 +114,6 @@
                 addSEs [lemma2_2, lemma2_3, lemma2_4, lemma2_5]) 1);
 val lemma2 = result();
 
-Goalw AC_defs "!!n. [| WO1; 1 le n |] ==> AC10(n)";
+Goalw AC_defs "[| WO1; 1 le n |] ==> AC10(n)";
 by (fast_tac (claset() addSIs [lemma1] addSEs [lemma2]) 1);
 qed "WO1_AC10";
--- a/src/ZF/AC/WO1_WO6.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/WO1_WO6.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -13,7 +13,7 @@
 
 (* ********************************************************************** *)
 
-Goalw WO_defs "!!Z. WO2 ==> WO3";
+Goalw WO_defs "WO2 ==> WO3";
 by (Fast_tac 1);
 qed "WO2_WO3";
 
@@ -32,20 +32,20 @@
 
 (* ********************************************************************** *)
 
-Goal "!!f. f: A->B ==> (lam x:A. {f`x}): A -> {{b}. b:B}";
+Goal "f: A->B ==> (lam x:A. {f`x}): A -> {{b}. b:B}";
 by (fast_tac (claset() addSIs [lam_type, apply_type]) 1);
 qed "lam_sets";
 
-Goalw [surj_def] "!!f. f:surj(A,B) ==> (UN a:A. {f`a}) = B";
+Goalw [surj_def] "f:surj(A,B) ==> (UN a:A. {f`a}) = B";
 by (fast_tac (claset() addSEs [apply_type]) 1);
 qed "surj_imp_eq_";
 
-Goal "!!f. [| f:surj(A,B); Ord(A) |] ==> (UN a<A. {f`a}) = B";
+Goal "[| f:surj(A,B); Ord(A) |] ==> (UN a<A. {f`a}) = B";
 by (fast_tac (claset() addSDs [surj_imp_eq_]
                 addSIs [ltI] addSEs [ltE]) 1);
 qed "surj_imp_eq";
 
-Goalw WO_defs "!!Z. WO1 ==> WO4(1)";
+Goalw WO_defs "WO1 ==> WO4(1)";
 by (rtac allI 1);
 by (eres_inst_tac [("x","A")] allE 1);
 by (etac exE 1);
@@ -61,20 +61,20 @@
 
 (* ********************************************************************** *)
 
-Goalw WO_defs "!!Z. [| m:nat; n:nat; m le n; WO4(m) |] ==> WO4(n)";
+Goalw WO_defs "[| m:nat; n:nat; m le n; WO4(m) |] ==> WO4(n)";
 by (fast_tac (claset() addIs [nat_le_imp_lepoll RSN (2, lepoll_trans)]) 1);
 qed "WO4_mono";
 
 (* ********************************************************************** *)
 
-Goalw WO_defs "!!Z. [| m:nat; 1 le m; WO4(m) |] ==> WO5";
+Goalw WO_defs "[| m:nat; 1 le m; WO4(m) |] ==> WO5";
     (*ZF_cs is essential: default claset's too slow*)
 by (fast_tac ZF_cs 1);
 qed "WO4_WO5";
 
 (* ********************************************************************** *)
 
-Goalw WO_defs "!!Z. WO5 ==> WO6";
+Goalw WO_defs "WO5 ==> WO6";
     (*ZF_cs is essential: default claset's too slow*)
 by (fast_tac ZF_cs 1);
 qed "WO5_WO6";
--- a/src/ZF/AC/WO1_WO7.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/WO1_WO7.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -19,7 +19,7 @@
 (* It is also easy to show that LEMMA implies WO1.                        *)
 (* ********************************************************************** *)
 
-Goalw [WO1_def] "!!Z. ALL X. ~Finite(X) -->  \
+Goalw [WO1_def] "ALL X. ~Finite(X) -->  \
 \               (EX R. well_ord(X,R) & ~well_ord(X,converse(R))) ==> WO1";
 by (rtac allI 1);
 by (etac allE 1);
@@ -37,7 +37,7 @@
 (* This statement can be proved by the following two theorems.            *)
 (* But moreover we need to show similar property for any well ordered     *)
 (* infinite set. It is not very difficult thanks to Isabelle order types  *)
-(* We show that if a set is well ordered by some relation and by it's     *)
+(* We show that if a set is well ordered by some relation and by its     *)
 (* converse, then apropriate order type is well ordered by the converse   *)
 (* of it's membership relation, which in connection with the previous     *)
 (* gives the conclusion.                                                  *)
@@ -51,10 +51,7 @@
 by (eres_inst_tac [("x","nat")] allE 1);
 by (etac disjE 1);
 by (fast_tac (claset() addSDs [nat_0I RSN (2,equals0D)]) 1);
-by (etac bexE 1);
-by (eres_inst_tac [("x","succ(x)")] allE 1);
-by (fast_tac (claset() addSIs [nat_succI, converseI, MemrelI, 
-                            nat_succI RSN (2, subsetD)]) 1);
+by (Blast_tac 1);
 qed "converse_Memrel_not_wf_on";
 
 Goalw [well_ord_def] 
@@ -62,7 +59,7 @@
 by (fast_tac (claset() addSDs [converse_Memrel_not_wf_on]) 1);
 qed "converse_Memrel_not_well_ord";
 
-Goal "!!A. [| well_ord(A,r); well_ord(A,converse(r)) |]  \
+Goal "[| well_ord(A,r); well_ord(A,converse(r)) |]  \
 \       ==> well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r))))";
 by (rtac ([ordertype_ord_iso RS ord_iso_sym RS ord_iso_rvimage_eq, 
                 Memrel_type RS (subset_Int_iff RS iffD1)] 
@@ -74,7 +71,7 @@
         THEN (assume_tac 1));
 qed "well_ord_converse_Memrel";
 
-Goalw [WO1_def] "!!Z. WO1 ==> ALL X. ~Finite(X) -->  \
+Goalw [WO1_def] "WO1 ==> ALL X. ~Finite(X) -->  \
 \                       (EX R. well_ord(X,R) & ~well_ord(X,converse(R)))";
 by (REPEAT (resolve_tac [allI,impI] 1));
 by (REPEAT (eresolve_tac [allE,exE] 1));
--- a/src/ZF/AC/WO1_WO8.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/WO1_WO8.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -9,7 +9,7 @@
 (* Trivial implication "WO1 ==> WO8"                                      *)
 (* ********************************************************************** *)
 
-Goalw WO_defs "!!Z. WO1 ==> WO8";
+Goalw WO_defs "WO1 ==> WO8";
 by (Fast_tac 1);
 qed "WO1_WO8";
 
@@ -17,7 +17,7 @@
 (* The proof of "WO8 ==> WO1" - faithful image of Rubin & Rubin's proof   *)
 (* ********************************************************************** *)
 
-Goalw WO_defs "!!Z. WO8 ==> WO1";
+Goalw WO_defs "WO8 ==> WO1";
 by (rtac allI 1);
 by (eres_inst_tac [("x","{{x}. x:A}")] allE 1);
 by (etac impE 1);
--- a/src/ZF/AC/WO2_AC16.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/WO2_AC16.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -18,7 +18,8 @@
 (* case of limit ordinal                                                  *)
 (* ********************************************************************** *)
 
-Goal "!!Z. [| ALL y<x. ALL z<a. z<y | (EX Y: F(y). f(z)<=Y)  \
+
+Goal "[| ALL y<x. ALL z<a. z<y | (EX Y: F(y). f(z)<=Y)  \
 \               --> (EX! Y. Y:F(y) & f(z)<=Y);  \
 \               ALL i j. i le j --> F(i) <= F(j); j le i; i<x; z<a;  \
 \               V: F(i); f(z)<=V; W:F(j); f(z)<=W |]  \
@@ -31,7 +32,8 @@
 by (etac ex1_two_eq 1 THEN (REPEAT (ares_tac [conjI] 1)));
 val lemma3_1 = result();
 
-Goal "!!Z. [| ALL y<x. ALL z<a. z<y | (EX Y: F(y). f(z)<=Y)  \
+
+Goal "[| ALL y<x. ALL z<a. z<y | (EX Y: F(y). f(z)<=Y)  \
 \               --> (EX! Y. Y:F(y) & f(z)<=Y);  \
 \               ALL i j. i le j --> F(i) <= F(j); i<x; j<x; z<a;  \
 \               V: F(i); f(z)<=V; W:F(j); f(z)<=W |]  \
@@ -42,7 +44,8 @@
 by (etac lemma3_1 1 THEN (REPEAT (assume_tac 1)));
 val lemma3 = result();
 
-Goal "!!a. [| ALL y<x. F(y) <= X &  \
+
+Goal "[| ALL y<x. F(y) <= X &  \
 \               (ALL x<a. x < y | (EX Y:F(y). fa(x) <= Y) -->  \
 \                       (EX! Y. Y : F(y) & fa(x) <= Y)); x < a |]  \
 \               ==> ALL y<x. ALL z<a. z < y | (EX Y:F(y). fa(z) <= Y) -->  \
@@ -52,7 +55,8 @@
 by (fast_tac (FOL_cs addSEs [oallE]) 1);
 val lemma4 = result();
 
-Goal "!!a. [| ALL y<x. F(y) <= X &  \
+
+Goal "[| ALL y<x. F(y) <= X &  \
 \               (ALL x<a. x < y | (EX Y:F(y). fa(x) <= Y) -->  \
 \                       (EX! Y. Y : F(y) & fa(x) <= Y)); \
 \               x < a; Limit(x); ALL i j. i le j --> F(i) <= F(j) |]  \
@@ -101,7 +105,8 @@
 (* dbl_Diff_eqpoll_Card                                                   *)
 (* ********************************************************************** *)
 
-Goal "!!A. [| A eqpoll a; Card(a); ~Finite(a); B lesspoll a;  \
+
+Goal "[| A eqpoll a; Card(a); ~Finite(a); B lesspoll a;  \
 \       C lesspoll a |] ==> A - B - C eqpoll a";
 by (rtac Diff_lesspoll_eqpoll_Card 1 THEN (REPEAT (assume_tac 1)));
 by (rtac Diff_lesspoll_eqpoll_Card 1 THEN (REPEAT (assume_tac 1)));
@@ -111,6 +116,7 @@
 (* Case of finite ordinals                                                *)
 (* ********************************************************************** *)
 
+
 Goalw [lesspoll_def]
         "!!X. [| Finite(X); ~Finite(a); Ord(a) |] ==> X lesspoll a";
 by (rtac conjI 1);
@@ -123,11 +129,13 @@
         subset_imp_lepoll RSN (2, eqpoll_imp_lepoll RS lepoll_trans)]) 1);
 qed "Finite_lesspoll_infinite_Ord";
 
-Goal "!!x. x:X ==> Union(X) = Union(X-{x}) Un x";
+
+Goal "x:X ==> Union(X) = Union(X-{x}) Un x";
 by (Fast_tac 1);
 qed "Union_eq_Un_Diff";
 
-Goal "!!n. n:nat ==> ALL X. X eqpoll n --> (ALL x:X. Finite(x))  \
+
+Goal "n:nat ==> ALL X. X eqpoll n --> (ALL x:X. Finite(x))  \
 \       --> Finite(Union(X))";
 by (etac nat_induct 1);
 by (fast_tac (claset() addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]
@@ -141,19 +149,22 @@
 by (fast_tac (claset() addSIs [Diff_sing_eqpoll]) 1);
 qed "Finite_Union_lemma";
 
-Goal "!!X. [| ALL x:X. Finite(x); Finite(X) |] ==> Finite(Union(X))";
+
+Goal "[| ALL x:X. Finite(x); Finite(X) |] ==> Finite(Union(X))";
 by (eresolve_tac [Finite_def RS def_imp_iff RS iffD1 RS bexE] 1);
 by (dtac Finite_Union_lemma 1);
 by (Fast_tac 1);
 qed "Finite_Union";
 
-Goalw [Finite_def] "!!x. [| x lepoll n; n:nat |] ==> Finite(x)";
+
+Goalw [Finite_def] "[| x lepoll n; n:nat |] ==> Finite(x)";
 by (fast_tac (claset()
         addEs [nat_into_Ord RSN (2, lepoll_imp_ex_le_eqpoll) RS exE,
         Ord_nat RSN (2, ltI) RSN (2, lt_trans1) RS ltD]) 1);
 qed "lepoll_nat_num_imp_Finite";
 
-Goal "!!X. [| ALL x:X. x lepoll n & x <= T; well_ord(T, R); X lepoll b;  \
+
+Goal "[| ALL x:X. x lepoll n & x <= T; well_ord(T, R); X lepoll b;  \
 \       b<a; ~Finite(a); Card(a); n:nat |]  \
 \       ==> Union(X) lesspoll a";
 by (excluded_middle_tac "Finite(X)" 1);
@@ -181,44 +192,51 @@
 (* recfunAC16_lepoll_index                                                *)
 (* ********************************************************************** *)
 
+
 Goal "A Un {a} = cons(a, A)";
 by (Fast_tac 1);
 qed "Un_sing_eq_cons";
 
-Goal "!!A. A lepoll B ==> A Un {a} lepoll succ(B)";
+
+Goal "A lepoll B ==> A Un {a} lepoll succ(B)";
 by (asm_simp_tac (simpset() addsimps [Un_sing_eq_cons, succ_def]) 1);
 by (eresolve_tac [mem_not_refl RSN (2, cons_lepoll_cong)] 1);
 qed "Un_lepoll_succ";
 
-Goal "!!a. Ord(a) ==> F(a) - (UN b<succ(a). F(b)) = 0";
+
+Goal "Ord(a) ==> F(a) - (UN b<succ(a). F(b)) = 0";
 by (fast_tac (claset() addSIs [OUN_I, le_refl]) 1);
 qed "Diff_UN_succ_empty";
 
-Goal "!!a. Ord(a) ==> F(a) Un X - (UN b<succ(a). F(b)) <= X";
+
+Goal "Ord(a) ==> F(a) Un X - (UN b<succ(a). F(b)) <= X";
 by (fast_tac (claset() addSIs [OUN_I, le_refl]) 1);
 qed "Diff_UN_succ_subset";
 
-Goal "!!x. Ord(x) ==>  \
+
+Goal "Ord(x) ==>  \
 \       recfunAC16(f, g, x, a) - (UN i<x. recfunAC16(f, g, i, a)) lepoll 1";
 by (etac Ord_cases 1);
 by (asm_simp_tac (simpset() addsimps [recfunAC16_0,
-                empty_subsetI RS subset_imp_lepoll]) 1);
-by (asm_simp_tac (simpset() addsimps [recfunAC16_Limit,
-                Diff_cancel, empty_subsetI RS subset_imp_lepoll]) 2);
+				      empty_subsetI RS subset_imp_lepoll]) 1);
+by (asm_simp_tac (simpset() addsimps [recfunAC16_Limit, Diff_cancel, 
+				      empty_subsetI RS subset_imp_lepoll]) 2);
 by (asm_simp_tac (simpset() addsimps [recfunAC16_succ]) 1);
-by (resolve_tac [conjI RS (split_if RS iffD2)] 1);
+by (rtac conjI 1);
 by (fast_tac (claset() addSIs [empty_subsetI RS subset_imp_lepoll]
-                addSEs [Diff_UN_succ_empty RS ssubst]) 1);
+                      addSEs [Diff_UN_succ_empty RS ssubst]) 1);
 by (fast_tac (claset() addSEs [Diff_UN_succ_subset RS subset_imp_lepoll RS
         (singleton_eqpoll_1 RS eqpoll_imp_lepoll RSN (2, lepoll_trans))]) 1);
 qed "recfunAC16_Diff_lepoll_1";
 
-Goal "!!z. [| z : F(x); Ord(x) |]  \
+
+Goal "[| z : F(x); Ord(x) |]  \
 \       ==> z:F(LEAST i. z:F(i)) - (UN j<(LEAST i. z:F(i)). F(j))";
 by (fast_tac (claset() addEs [less_LeastE] addSEs [OUN_E, LeastI]) 1);
 qed "in_Least_Diff";
 
-Goal "!!w. [| (LEAST i. w:F(i)) = (LEAST i. z:F(i));  \
+
+Goal "[| (LEAST i. w:F(i)) = (LEAST i. z:F(i));  \
 \       w:(UN i<a. F(i)); z:(UN i<a. F(i)) |]  \
 \       ==> EX b<a. w:(F(b) - (UN c<b. F(c))) & z:(F(b) - (UN c<b. F(c)))";
 by (REPEAT (etac OUN_E 1));
@@ -231,11 +249,13 @@
         THEN (REPEAT (assume_tac 1)));
 qed "Least_eq_imp_ex";
 
-Goal "!!A. [| A lepoll 1; a:A; b:A |] ==> a=b";
+
+Goal "[| A lepoll 1; a:A; b:A |] ==> a=b";
 by (fast_tac (claset() addSDs [lepoll_1_is_sing]) 1);
 qed "two_in_lepoll_1";
 
-Goal "!!a. [| ALL i<a. F(i)-(UN j<i. F(j)) lepoll 1; Limit(a) |]  \
+
+Goal "[| ALL i<a. F(i)-(UN j<i. F(j)) lepoll 1; Limit(a) |]  \
 \       ==> (UN x<a. F(x)) lepoll a";
 by (resolve_tac [lepoll_def RS (def_imp_iff RS iffD2)] 1);
 by (res_inst_tac [("x","lam z: (UN x<a. F(x)). LEAST i. z:F(i)")] exI 1);
@@ -254,31 +274,35 @@
 by (fast_tac (claset() addSEs [two_in_lepoll_1]) 1);
 qed "UN_lepoll_index";
 
-Goal "!!y. Ord(y) ==> recfunAC16(f, fa, y, a) lepoll y";
+
+Goal "Ord(y) ==> recfunAC16(f, fa, y, a) lepoll y";
 by (etac trans_induct 1);
 by (etac Ord_cases 1);
 by (asm_simp_tac (simpset() addsimps [recfunAC16_0, lepoll_refl]) 1);
 by (asm_simp_tac (simpset() addsimps [recfunAC16_succ]) 1);
-by (fast_tac (claset() addIs [conjI RS (split_if RS iffD2)]
+by (fast_tac (claset() 
         addSDs [succI1 RSN (2, bspec)]
         addSEs [subset_succI RS subset_imp_lepoll RSN (2, lepoll_trans),
                 Un_lepoll_succ]) 1);
 by (asm_simp_tac (simpset() addsimps [recfunAC16_Limit]) 1);
 by (fast_tac (claset() addSEs [lt_Ord RS recfunAC16_Diff_lepoll_1]
-        addSIs [UN_lepoll_index]) 1);
+                       addSIs [UN_lepoll_index]) 1);
 qed "recfunAC16_lepoll_index";
 
-Goal "!!y. [| recfunAC16(f,g,y,a) <= {X: Pow(A). X eqpoll n};  \
+
+Goal "[| recfunAC16(f,g,y,a) <= {X: Pow(A). X eqpoll n};  \
 \               A eqpoll a; y<a; ~Finite(a); Card(a); n:nat |]  \
 \               ==> Union(recfunAC16(f,g,y,a)) lesspoll a";
 by (eresolve_tac [eqpoll_def RS def_imp_iff RS iffD1 RS exE] 1);
 by (rtac Union_lesspoll 1 THEN (TRYALL assume_tac));
 by (eresolve_tac [lt_Ord RS recfunAC16_lepoll_index] 3);
 by (eresolve_tac [[bij_is_inj, Card_is_Ord RS well_ord_Memrel] MRS
-        well_ord_rvimage] 2 THEN (assume_tac 2));
+		  well_ord_rvimage] 2 
+    THEN (assume_tac 2));
 by (fast_tac (claset() addSEs [eqpoll_imp_lepoll]) 1);
 qed "Union_recfunAC16_lesspoll";
 
+
 Goal
   "!!a. [| recfunAC16(f, fa, y, a) <= {X: Pow(A) . X eqpoll succ(k #+ m)};  \
 \       Card(a); ~ Finite(a); A eqpoll a;  \
@@ -302,7 +326,8 @@
                              nat_sum_eqpoll_sum RSN (3, 
                              eqpoll_trans RS eqpoll_trans))) |> standard;
 
-Goal "!!x. [| x : Pow(A - B - fa`i); x eqpoll m;  \
+
+Goal "[| x : Pow(A - B - fa`i); x eqpoll m;  \
 \       fa : bij(a, {x: Pow(A) . x eqpoll k}); i<a; k:nat; m:nat |]  \
 \       ==> fa ` i Un x : {x: Pow(A) . x eqpoll k #+ m}";
 by (rtac CollectI 1);
@@ -319,7 +344,8 @@
 (* Lemmas simplifying assumptions                                         *)
 (* ********************************************************************** *)
 
-Goal "!!j. [| ALL y<succ(j). F(y)<=X & (ALL x<a. x<y | P(x,y)  \
+
+Goal "[| ALL y<succ(j). F(y)<=X & (ALL x<a. x<y | P(x,y)  \
 \       --> Q(x,y)); succ(j)<a |]  \
 \       ==> F(j)<=X & (ALL x<a. x<j | P(x,j) --> Q(x,j))";
 by (dtac ospec 1);
@@ -327,8 +353,9 @@
         THEN (REPEAT (assume_tac 1)));
 val lemma6 = result();
 
-Goal "!!j. [| F(j)<=X; (ALL x<a. x<j | P(x,j) --> Q(x,j)); succ(j)<a |]  \
-\       ==> P(j,j) --> F(j) <= X & (ALL x<a. x le j | P(x,j) --> Q(x,j))";
+
+Goal "[| ALL x<a. x<j | P(x,j) --> Q(x,j);  succ(j)<a |]  \
+\     ==> P(j,j) --> (ALL x<a. x le j | P(x,j) --> Q(x,j))";
 by (fast_tac (claset() addSEs [leE]) 1);
 val lemma7 = result();
 
@@ -337,7 +364,8 @@
 (* ordinal there is a set satisfying certain properties                   *)
 (* ********************************************************************** *)
 
-Goal "!!A. [| A eqpoll a; ~ Finite(a); Ord(a); m:nat |]  \
+
+Goal "[| A eqpoll a; ~ Finite(a); Ord(a); m:nat |]  \
 \       ==> EX X:Pow(A). X eqpoll m";
 by (eresolve_tac [Ord_nat RSN (2, ltI) RS 
                 (nat_le_infinite_Ord RSN (2, lt_trans2)) RS 
@@ -348,11 +376,13 @@
 by (fast_tac (claset() addSEs [eqpoll_sym]) 1);
 qed "ex_subset_eqpoll";
 
-Goal "!!A. [| A <= B Un C; A Int C = 0 |] ==> A <= B";
+
+Goal "[| A <= B Un C; A Int C = 0 |] ==> A <= B";
 by (fast_tac (claset() addDs [equals0D]) 1);
 qed "subset_Un_disjoint";
 
-Goal "!!F. [| X:Pow(A - Union(B) -C); T:B; F<=T |] ==> F Int X = 0";
+
+Goal "[| X:Pow(A - Union(B) -C); T:B; F<=T |] ==> F Int X = 0";
 by (fast_tac (claset() addSIs [equals0I]) 1);
 qed "Int_empty";
 
@@ -360,11 +390,13 @@
 (* equipollent subset (and finite) is the whole set                       *)
 (* ********************************************************************** *)
 
-Goal "!!A. [| A <= B; a : A; A - {a} = B - {a} |] ==> A = B";
+
+Goal "[| A <= B; a : A; A - {a} = B - {a} |] ==> A = B";
 by (fast_tac (claset() addSEs [equalityE]) 1);
 qed "Diffs_eq_imp_eq";
 
-Goal "!!A. m:nat ==> ALL A B. A <= B & m lepoll A & B lepoll m --> A=B";
+
+Goal "m:nat ==> ALL A B. A <= B & m lepoll A & B lepoll m --> A=B";
 by (etac nat_induct 1);
 by (fast_tac (claset() addSDs [lepoll_0_is_0]) 1);
 by (REPEAT (resolve_tac [allI, impI] 1));
@@ -382,11 +414,13 @@
         THEN REPEAT (assume_tac 1));
 qed "subset_imp_eq_lemma";
 
-Goal "!!A. [| A <= B; m lepoll A; B lepoll m; m:nat |] ==> A=B";
+
+Goal "[| A <= B; m lepoll A; B lepoll m; m:nat |] ==> A=B";
 by (fast_tac (FOL_cs addSDs [subset_imp_eq_lemma]) 1);
 qed "subset_imp_eq";
 
-Goal "!!f. [| f:bij(a, {Y:X. Y eqpoll succ(k)}); k:nat; f`b<=f`y; b<a;  \
+
+Goal "[| f:bij(a, {Y:X. Y eqpoll succ(k)}); k:nat; f`b<=f`y; b<a;  \
 \       y<a |] ==> b=y";
 by (dtac subset_imp_eq 1);
 by (etac nat_succI 3);
@@ -398,6 +432,7 @@
 by (fast_tac (claset() addSDs [ltD]) 1);
 qed "bij_imp_arg_eq";
 
+
 Goal
   "!!a. [| recfunAC16(f, fa, y, a) <= {X: Pow(A) . X eqpoll succ(k #+ m)};  \
 \       Card(a); ~ Finite(a); A eqpoll a;  \
@@ -434,6 +469,7 @@
 (* ordinal there is a number of the set satisfying certain properties     *)
 (* ********************************************************************** *)
 
+
 Goal
   "!!a. [| recfunAC16(f, fa, y, a) <= {X: Pow(A) . X eqpoll succ(k #+ m)};  \
 \       Card(a); ~ Finite(a); A eqpoll a;  \
@@ -452,7 +488,8 @@
         THEN REPEAT (ares_tac [Card_is_Ord] 1));
 qed "ex_next_Ord";
 
-Goal "!!Z. [| EX! Y. Y:Z & P(Y); ~P(W) |] ==> EX! Y. Y: (Z Un {W}) & P(Y)";
+
+Goal "[| EX! Y. Y:Z & P(Y); ~P(W) |] ==> EX! Y. Y: (Z Un {W}) & P(Y)";
 by (Fast_tac 1);
 qed "ex1_in_Un_sing";
 
@@ -460,7 +497,8 @@
 (* Lemma simplifying assumptions                                          *)
 (* ********************************************************************** *)
 
-Goal "!!j. [| ALL x<a. x<j | (EX xa:F(j). P(x, xa))  \
+
+Goal "[| ALL x<a. x<j | (EX xa:F(j). P(x, xa))  \
 \       --> (EX! Y. Y:F(j) & P(x, Y)); F(j) <= X;  \
 \       L : X; P(j, L) & (ALL x<a. P(x, L) --> (ALL xa:F(j). ~P(x, xa))) |]  \
 \       ==> F(j) Un {L} <= X &  \
@@ -470,18 +508,7 @@
 by (fast_tac (claset() addSIs [singleton_subsetI]) 1);
 by (rtac oallI 1);
 by (etac oallE 1 THEN (contr_tac 2));
-by (rtac impI 1);
-by (etac disjE 1);
-by (etac leE 1);
-by (eresolve_tac [disjI1 RSN (2, impE)] 1 THEN (assume_tac 1));
-by (rtac ex1E 1 THEN (assume_tac 1));
-by (etac ex1_in_Un_sing 1);
-by (Fast_tac 1);
-by (Deepen_tac 2 1);
-by (etac bexE 1);
-by (etac UnE 1);
-by (fast_tac (claset() delrules [ex_ex1I] addSIs [ex1_in_Un_sing]) 1);
-by (Deepen_tac 2 1);
+by (blast_tac (claset() addSEs [leE]) 1);
 val lemma8 = result();
 
 (* ********************************************************************** *)
@@ -489,8 +516,8 @@
 (* lemma main_induct                                                      *)
 (* ********************************************************************** *)
 
-Goal
-        "!!a. [| b < a; f : bij(a, {Y: Pow(A) . Y eqpoll succ(k #+ m)});  \
+
+Goal "[| b < a; f : bij(a, {Y: Pow(A) . Y eqpoll succ(k #+ m)});  \
 \       fa : bij(a, {Y: Pow(A) . Y eqpoll succ(k)});  \
 \       ~Finite(a); Card(a); A eqpoll a; k : nat; m : nat |]  \
 \       ==> recfunAC16(f, fa, b, a) <= {X: Pow(A) . X eqpoll succ(k #+ m)} &  \
@@ -508,17 +535,19 @@
 (* case succ *)
 by (hyp_subst_tac 1);
 by (eresolve_tac [lemma6 RS conjE] 1 THEN (assume_tac 1));
-by (asm_simp_tac (simpset() addsimps [recfunAC16_succ]) 1);
+by (asm_simp_tac (simpset() delsplits [split_if]
+			    addsimps [recfunAC16_succ]) 1);
 by (resolve_tac [conjI RS (split_if RS iffD2)] 1);
-by (etac lemma7 1 THEN (REPEAT (assume_tac 1)));
+by (Asm_simp_tac 1);
+by (etac lemma7 1 THEN  assume_tac 1);
 by (rtac impI 1);
 by (resolve_tac [ex_next_Ord RS oexE] 1 
-        THEN REPEAT (ares_tac [le_refl RS lt_trans] 1));
+    THEN REPEAT (ares_tac [le_refl RS lt_trans] 1));
 by (etac lemma8 1 THEN (assume_tac 1));
 by (resolve_tac [bij_is_fun RS apply_type] 1 THEN (assume_tac 1));
 by (eresolve_tac [Least_le RS lt_trans2 RS ltD] 1 
         THEN REPEAT (ares_tac [lt_Ord, succ_leI] 1));
-(*VERY SLOW!  24 secs??*)
+(*SLOW!  5 secs?*)
 by (eresolve_tac [lt_Ord RSN (2, LeastI)] 1 THEN (assume_tac 1));
 qed "main_induct";
 
@@ -574,6 +603,7 @@
 (* The target theorem                                                     *)
 (* ********************************************************************** *)
 
+
 Goalw [AC16_def]
         "!!n k. [| WO2; 0<m; k:nat; m:nat |] ==> AC16(k #+ m,k)";
 by (rtac allI 1);
--- a/src/ZF/AC/WO6_WO1.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/WO6_WO1.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -36,7 +36,7 @@
 by (Blast_tac 1);
 qed "domain_uu_subset";
 
-Goal "!! a. ALL b<a. f`b lepoll m ==> \
+Goal "ALL b<a. f`b lepoll m ==> \
 \               ALL b<a. ALL g<a. ALL d<a. domain(uu(f,b,g,d)) lepoll m";
 by (fast_tac (claset() addSEs
         [domain_uu_subset RS subset_imp_lepoll RS lepoll_trans]) 1);
@@ -50,7 +50,7 @@
 by (Blast_tac 1);
 qed "uu_subset2";
 
-Goal "!! a. [| ALL b<a. f`b lepoll m;  d<a |] ==> uu(f,b,g,d) lepoll m";
+Goal "[| ALL b<a. f`b lepoll m;  d<a |] ==> uu(f,b,g,d) lepoll m";
 by (fast_tac (claset()
         addSEs [uu_subset2 RS subset_imp_lepoll RS lepoll_trans]) 1);
 qed "uu_lepoll_m";
@@ -71,7 +71,7 @@
 (* ********************************************************************** *)
 (* Lemmas used in both cases                                              *)
 (* ********************************************************************** *)
-Goal "!!a C. Ord(a) ==> (UN b<a++a. C(b)) = (UN b<a. C(b) Un C(a++b))";
+Goal "Ord(a) ==> (UN b<a++a. C(b)) = (UN b<a. C(b) Un C(a++b))";
 by (fast_tac (claset() addSIs [equalityI] addIs [ltI] 
                     addSDs [lt_oadd_disj]
                     addSEs [lt_oadd1, oadd_lt_mono2]) 1);
@@ -84,7 +84,6 @@
 
 Goalw [vv1_def] "vv1(f,m,b) <= f`b";
 by (rtac (LetI RS LetI) 1);
-by (split_tac [split_if] 1);
 by (simp_tac (simpset() addsimps [domain_uu_subset]) 1);
 qed "vv1_subset";
 
@@ -108,7 +107,7 @@
 (* ********************************************************************** *)
 (* every value of defined function is less than or equipollent to m       *)
 (* ********************************************************************** *)
-Goal "!!a b. [| P(a, b);  Ord(a);  Ord(b);  \
+Goal "[| P(a, b);  Ord(a);  Ord(b);  \
 \               Least_a = (LEAST a. EX x. Ord(x) & P(a, x)) |]  \
 \               ==> P(Least_a, LEAST b. P(Least_a, b))";
 by (etac ssubst 1);
@@ -132,8 +131,7 @@
 by (Asm_simp_tac 1);
 by (safe_tac (claset() addSEs [lt_oadd_odiff_cases]));
 (*Case b<a   : show vv1(f,m,b) lepoll m *)
-by (asm_simp_tac (simpset() addsimps [vv1_def, Let_def] 
-                        addsplits [split_if]) 1);
+by (asm_simp_tac (simpset() addsimps [vv1_def, Let_def]) 1);
 by (fast_tac (claset() addIs [nested_Least_instance RS conjunct2]
                 addSEs [lt_Ord]
                 addSIs [empty_lepollI]) 1);
@@ -159,7 +157,7 @@
 (* Case 2 : vv2_subset                                                    *)
 (* ********************************************************************** *)
 
-Goalw [uu_def] "!!f. [| b<a;  g<a;  f`b~=0;  f`g~=0;        \
+Goalw [uu_def] "[| b<a;  g<a;  f`b~=0;  f`g~=0;        \
 \                           y*y <= y;  (UN b<a. f`b)=y          \
 \                        |] ==> EX d<a. uu(f,b,g,d) ~= 0";
 by (fast_tac (claset() addSIs [not_emptyI] 
@@ -167,7 +165,7 @@
                     addSEs [not_emptyE]) 1);
 qed "ex_d_uu_not_empty";
 
-Goal "!!f. [| b<a; g<a; f`b~=0; f`g~=0;  \
+Goal "[| b<a; g<a; f`b~=0; f`g~=0;  \
 \                       y*y<=y; (UN b<a. f`b)=y |]  \
 \               ==> uu(f,b,g,LEAST d. (uu(f,b,g,d) ~= 0)) ~= 0";
 by (dtac ex_d_uu_not_empty 1 THEN REPEAT (assume_tac 1));
@@ -179,7 +177,7 @@
                 sym RSN (2, subst_elem) RS domainI RS not_emptyI] 1));
 qed "not_empty_rel_imp_domain";
 
-Goal "!!f. [| b<a; g<a; f`b~=0; f`g~=0;  \
+Goal "[| b<a; g<a; f`b~=0; f`g~=0;  \
 \                       y*y <= y; (UN b<a. f`b)=y |]  \
 \               ==> (LEAST d. uu(f,b,g,d) ~= 0) < a";
 by (eresolve_tac [ex_d_uu_not_empty RS oexE] 1
@@ -193,7 +191,7 @@
 qed "subset_Diff_sing";
 
 (*Could this be proved more directly?*)
-Goal "!!A B. [| A lepoll m; m lepoll B; B <= A; m:nat |] ==> A=B";
+Goal "[| A lepoll m; m lepoll B; B <= A; m:nat |] ==> A=B";
 by (etac natE 1);
 by (fast_tac (claset() addSDs [lepoll_0_is_0] addSIs [equalityI]) 1);
 by (hyp_subst_tac 1);
@@ -222,7 +220,7 @@
         uu_subset1 RSN (4, rel_is_fun)))] 1
         THEN TRYALL assume_tac);
 by (rtac (eqpoll_sym RS eqpoll_imp_lepoll RSN (2, supset_lepoll_imp_eq)) 1);
-by (REPEAT (fast_tac (claset() addSIs [domain_uu_subset, nat_succI]) 1));
+by (REPEAT (fast_tac (claset() addSIs [domain_uu_subset]) 1));
 qed "uu_Least_is_fun";
 
 Goalw [vv2_def]
@@ -264,8 +262,7 @@
 
 Goalw [vv2_def]
     "!!m. [| m:nat; m~=0 |] ==> vv2(f,b,g,s) lepoll m";
-by (asm_simp_tac (simpset() addsimps [empty_lepollI]
-                              addsplits [split_if]) 1);
+by (asm_simp_tac (simpset() addsimps [empty_lepollI]) 1);
 by (fast_tac (claset()
         addSDs [le_imp_subset RS subset_imp_lepoll RS lepoll_0_is_0]
         addSIs [singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans,
@@ -279,9 +276,8 @@
 by (excluded_middle_tac "f`g = 0" 1);
 by (asm_simp_tac (simpset() addsimps [empty_lepollI]) 2);
 by (dtac ospec 1 THEN (assume_tac 1));
-by (rtac Diff_lepoll 1
-        THEN (TRYALL assume_tac));
-by (asm_simp_tac (simpset() addsimps [vv2_def, split_if, not_emptyI]) 1);
+by (rtac Diff_lepoll 1 THEN (TRYALL assume_tac));
+by (asm_simp_tac (simpset() addsimps [vv2_def, not_emptyI]) 1);
 qed "ww2_lepoll";
 
 Goalw [gg2_def]
@@ -336,14 +332,14 @@
 by (fast_tac (claset() addIs [rec_succ RS ssubst]) 1);
 qed "z_n_subset_z_succ_n";
 
-Goal "!!n. [| ALL n:nat. f(n)<=f(succ(n)); n le m; n : nat; m: nat |]  \
+Goal "[| ALL n:nat. f(n)<=f(succ(n)); n le m; n : nat; m: nat |]  \
 \              ==> f(n)<=f(m)";
 by (eres_inst_tac [("P","n le m")] rev_mp 1);
 by (res_inst_tac [("P","%z. n le z --> f(n) <= f(z)")] nat_induct 1);
 by (REPEAT (fast_tac le_cs 1));
 qed "le_subsets";
 
-Goal "!!n m. [| n le m; m:nat |] ==>  \
+Goal "[| n le m; m:nat |] ==>  \
 \       rec(n, x, %k r. r Un r*r) <= rec(m, x, %k r. r Un r*r)";
 by (resolve_tac [z_n_subset_z_succ_n RS le_subsets] 1 
     THEN (TRYALL assume_tac));
@@ -386,18 +382,18 @@
 (*      1 : NN(y) ==> y can be well-ordered                               *)
 (* ********************************************************************** *)
 
-Goal "!!f. [| (UN b<a. f`b)=y; x:y; ALL b<a. f`b lepoll 1; Ord(a) |]  \
+Goal "[| (UN b<a. f`b)=y; x:y; ALL b<a. f`b lepoll 1; Ord(a) |]  \
 \               ==> EX c<a. f`c = {x}";
 by (fast_tac (claset() addSEs [lepoll_1_is_sing]) 1);
 val lemma1 = result();
 
-Goal "!!f. [| (UN b<a. f`b)=y; x:y; ALL b<a. f`b lepoll 1; Ord(a) |]  \
+Goal "[| (UN b<a. f`b)=y; x:y; ALL b<a. f`b lepoll 1; Ord(a) |]  \
 \               ==> f` (LEAST i. f`i = {x}) = {x}";
 by (dtac lemma1 1 THEN REPEAT (assume_tac 1));
 by (fast_tac (claset() addSEs [lt_Ord] addIs [LeastI]) 1);
 val lemma2 = result();
 
-Goalw [NN_def] "!!y. 1 : NN(y) ==> EX a f. Ord(a) & f:inj(y, a)";
+Goalw [NN_def] "1 : NN(y) ==> EX a f. Ord(a) & f:inj(y, a)";
 by (etac CollectE 1);
 by (REPEAT (eresolve_tac [exE, conjE] 1));
 by (res_inst_tac [("x","a")] exI 1);
@@ -410,7 +406,7 @@
 by (fast_tac (claset() addSIs [the_equality]) 1);
 qed "NN_imp_ex_inj";
 
-Goal "!!y. [| y*y <= y; 1 : NN(y) |] ==> EX r. well_ord(y, r)";
+Goal "[| y*y <= y; 1 : NN(y) |] ==> EX r. well_ord(y, r)";
 by (dtac NN_imp_ex_inj 1);
 by (fast_tac (claset() addSEs [well_ord_Memrel RSN (2,  well_ord_rvimage)]) 1);
 qed "y_well_ord";
@@ -439,11 +435,11 @@
 by (REPEAT (ares_tac prems 1));
 qed "rev_induct";
 
-Goalw [NN_def] "!!n. n:NN(y) ==> n:nat";
+Goalw [NN_def] "n:NN(y) ==> n:nat";
 by (etac CollectD1 1);
 qed "NN_into_nat";
 
-Goal "!!n. [| n:NN(y); y*y <= y; n~=0 |] ==> 1:NN(y)";
+Goal "[| n:NN(y); y*y <= y; n~=0 |] ==> 1:NN(y)";
 by (rtac rev_induct 1 THEN REPEAT (ares_tac [NN_into_nat] 1));
 by (rtac lemma_ii 1 THEN REPEAT (assume_tac 1));
 val lemma3 = result();
@@ -453,12 +449,12 @@
 (* ********************************************************************** *)
 
 (* another helpful lemma *)
-Goalw [NN_def] "!!y. 0:NN(y) ==> y=0";
+Goalw [NN_def] "0:NN(y) ==> y=0";
 by (fast_tac (claset() addSIs [equalityI] 
                     addSDs [lepoll_0_is_0] addEs [subst]) 1);
 qed "NN_y_0";
 
-Goalw [WO1_def] "!!Z. WO6 ==> WO1";
+Goalw [WO1_def] "WO6 ==> WO1";
 by (rtac allI 1);
 by (excluded_middle_tac "A=0" 1);
 by (fast_tac (claset() addSIs [well_ord_Memrel, nat_0I RS nat_into_Ord]) 2);
--- a/src/ZF/AC/WO_AC.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/WO_AC.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -7,16 +7,16 @@
 
 open WO_AC;
 
-Goal "!!A. [| well_ord(Union(A),r); 0~:A; B:A |]  \
+Goal "[| well_ord(Union(A),r); 0~:A; B:A |]  \
 \               ==> (THE b. first(b,B,r)) : B";
 by (fast_tac (claset() addSEs [well_ord_imp_ex1_first RS theI RS
                 (first_def RS def_imp_iff RS iffD1 RS conjunct1)]) 1);
 qed "first_in_B";
 
-Goal "!!A. [| well_ord(Union(A), R); 0~:A |] ==> EX f. f:(PROD X:A. X)";
+Goal "[| well_ord(Union(A), R); 0~:A |] ==> EX f. f:(PROD X:A. X)";
 by (fast_tac (claset() addSEs [first_in_B] addSIs [lam_type]) 1);
 qed "ex_choice_fun";
 
-Goal "!!A. well_ord(A, R) ==> EX f. f:(PROD X: Pow(A)-{0}. X)";
+Goal "well_ord(A, R) ==> EX f. f:(PROD X: Pow(A)-{0}. X)";
 by (fast_tac (claset() addSEs [well_ord_subset RS ex_choice_fun]) 1);
 qed "ex_choice_fun_Pow";
--- a/src/ZF/AC/recfunAC16.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/AC/recfunAC16.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -23,7 +23,7 @@
 by (rtac transrec2_succ 1);
 qed "recfunAC16_succ";
 
-Goalw [recfunAC16_def] "!!i. Limit(i)  \
+Goalw [recfunAC16_def] "Limit(i)  \
 \       ==> recfunAC16(f,fa,i,a) = (UN j<i. recfunAC16(f,fa,j,a))";
 by (etac transrec2_Limit 1);
 qed "recfunAC16_Limit";
--- a/src/ZF/Arith.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Arith.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -47,7 +47,7 @@
 Addsimps [rec_type, nat_0_le, nat_le_refl];
 val nat_typechecks = [rec_type, nat_0I, nat_1I, nat_succI, Ord_nat];
 
-Goal "!!k. [| 0<k; k: nat |] ==> EX j: nat. k = succ(j)";
+Goal "[| 0<k; k: nat |] ==> EX j: nat. k = succ(j)";
 by (etac rev_mp 1);
 by (etac nat_induct 1);
 by (Simp_tac 1);
@@ -198,11 +198,11 @@
 
 Addsimps [mult_0_right, mult_succ_right];
 
-Goal "!!n. n:nat ==> 1 #* n = n";
+Goal "n:nat ==> 1 #* n = n";
 by (Asm_simp_tac 1);
 qed "mult_1";
 
-Goal "!!n. n:nat ==> n #* 1 = n";
+Goal "n:nat ==> n #* 1 = n";
 by (Asm_simp_tac 1);
 qed "mult_1_right";
 
@@ -255,7 +255,7 @@
     (ALLGOALS (asm_simp_tac (simpset() addsimps prems))) ]);
 
 (*Addition is the inverse of subtraction*)
-Goal "!!m n. [| n le m;  m:nat |] ==> n #+ (m#-n) = m";
+Goal "[| n le m;  m:nat |] ==> n #+ (m#-n) = m";
 by (forward_tac [lt_nat_in_nat] 1);
 by (etac nat_succI 1);
 by (etac rev_mp 1);
@@ -264,7 +264,7 @@
 qed "add_diff_inverse";
 
 (*Proof is IDENTICAL to that above*)
-Goal "!!m n. [| n le m;  m:nat |] ==> succ(m) #- n = succ(m#-n)";
+Goal "[| n le m;  m:nat |] ==> succ(m) #- n = succ(m#-n)";
 by (forward_tac [lt_nat_in_nat] 1);
 by (etac nat_succI 1);
 by (etac rev_mp 1);
@@ -280,20 +280,17 @@
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [mnat])));
 qed "diff_add_inverse";
 
-Goal
-    "!!m n. [| m:nat;  n:nat |] ==> (m#+n) #- n = m";
+Goal "[| m:nat;  n:nat |] ==> (m#+n) #- n = m";
 by (res_inst_tac [("m1","m")] (add_commute RS ssubst) 1);
 by (REPEAT (ares_tac [diff_add_inverse] 1));
 qed "diff_add_inverse2";
 
-Goal
-    "!!k. [| k:nat; m: nat; n: nat |] ==> (k#+m) #- (k#+n) = m #- n";
+Goal "[| k:nat; m: nat; n: nat |] ==> (k#+m) #- (k#+n) = m #- n";
 by (nat_ind_tac "k" [] 1);
 by (ALLGOALS Asm_simp_tac);
 qed "diff_cancel";
 
-Goal
-    "!!n. [| k:nat; m: nat; n: nat |] ==> (m#+k) #- (n#+k) = m #- n";
+Goal "[| k:nat; m: nat; n: nat |] ==> (m#+k) #- (n#+k) = m #- n";
 val add_commute_k = read_instantiate [("n","k")] add_commute;
 by (asm_simp_tac (simpset() addsimps [add_commute_k, diff_cancel]) 1);
 qed "diff_cancel2";
@@ -306,14 +303,12 @@
 
 (** Difference distributes over multiplication **)
 
-Goal 
-  "!!m n. [| m:nat; n: nat; k:nat |] ==> (m #- n) #* k = (m #* k) #- (n #* k)";
+Goal "[| m:nat; n: nat; k:nat |] ==> (m #- n) #* k = (m #* k) #- (n #* k)";
 by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_cancel])));
 qed "diff_mult_distrib" ;
 
-Goal 
-  "!!m. [| m:nat; n: nat; k:nat |] ==> k #* (m #- n) = (k #* m) #- (k #* n)";
+Goal "[| m:nat; n: nat; k:nat |] ==> k #* (m #- n) = (k #* m) #- (k #* n)";
 val mult_commute_k = read_instantiate [("m","k")] mult_commute;
 by (asm_simp_tac (simpset() addsimps 
                   [mult_commute_k, diff_mult_distrib]) 1);
@@ -321,7 +316,7 @@
 
 (*** Remainder ***)
 
-Goal "!!m n. [| 0<n;  n le m;  m:nat |] ==> m #- n < m";
+Goal "[| 0<n;  n le m;  m:nat |] ==> m #- n < m";
 by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1);
 by (etac rev_mp 1);
 by (etac rev_mp 1);
@@ -338,16 +333,16 @@
                                   not_lt_iff_le RS iffD2];
 
 (*Type checking depends upon termination!*)
-Goalw [mod_def] "!!m n. [| 0<n;  m:nat;  n:nat |] ==> m mod n : nat";
+Goalw [mod_def] "[| 0<n;  m:nat;  n:nat |] ==> m mod n : nat";
 by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1));
 qed "mod_type";
 
-Goal "!!m n. [| 0<n;  m<n |] ==> m mod n = m";
+Goal "[| 0<n;  m<n |] ==> m mod n = m";
 by (rtac (mod_def RS def_transrec RS trans) 1);
 by (asm_simp_tac div_ss 1);
 qed "mod_less";
 
-Goal "!!m n. [| 0<n;  n le m;  m:nat |] ==> m mod n = (m#-n) mod n";
+Goal "[| 0<n;  n le m;  m:nat |] ==> m mod n = (m#-n) mod n";
 by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1);
 by (rtac (mod_def RS def_transrec RS trans) 1);
 by (asm_simp_tac div_ss 1);
@@ -363,13 +358,12 @@
 by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1));
 qed "div_type";
 
-Goal "!!m n. [| 0<n;  m<n |] ==> m div n = 0";
+Goal "[| 0<n;  m<n |] ==> m div n = 0";
 by (rtac (div_def RS def_transrec RS trans) 1);
 by (asm_simp_tac div_ss 1);
 qed "div_less";
 
-Goal
- "!!m n. [| 0<n;  n le m;  m:nat |] ==> m div n = succ((m#-n) div n)";
+Goal "[| 0<n;  n le m;  m:nat |] ==> m div n = succ((m#-n) div n)";
 by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1);
 by (rtac (div_def RS def_transrec RS trans) 1);
 by (asm_simp_tac div_ss 1);
@@ -378,8 +372,7 @@
 Addsimps [div_type, div_less, div_geq];
 
 (*A key result*)
-Goal
-    "!!m n. [| 0<n;  m:nat;  n:nat |] ==> (m div n)#*n #+ m mod n = m";
+Goal "[| 0<n;  m:nat;  n:nat |] ==> (m div n)#*n #+ m mod n = m";
 by (etac complete_induct 1);
 by (excluded_middle_tac "x<n" 1);
 (*case x<n*)
@@ -392,8 +385,7 @@
 
 (*** Further facts about mod (mainly for mutilated checkerboard ***)
 
-Goal
-    "!!m n. [| 0<n;  m:nat;  n:nat |] ==> \
+Goal "[| 0<n;  m:nat;  n:nat |] ==> \
 \           succ(m) mod n = if(succ(m mod n) = n, 0, succ(m mod n))";
 by (etac complete_induct 1);
 by (excluded_middle_tac "succ(x)<n" 1);
@@ -410,7 +402,7 @@
 by (asm_simp_tac (simpset() addsimps [mod_less, diff_self_eq_0]) 1);
 qed "mod_succ";
 
-Goal "!!m n. [| 0<n;  m:nat;  n:nat |] ==> m mod n < n";
+Goal "[| 0<n;  m:nat;  n:nat |] ==> m mod n < n";
 by (etac complete_induct 1);
 by (excluded_middle_tac "x<n" 1);
 (*case x<n*)
@@ -422,22 +414,20 @@
 qed "mod_less_divisor";
 
 
-Goal
-    "!!k b. [| k: nat; b<2 |] ==> k mod 2 = b | k mod 2 = if(b=1,0,1)";
+Goal "[| k: nat; b<2 |] ==> k mod 2 = b | k mod 2 = if(b=1,0,1)";
 by (subgoal_tac "k mod 2: 2" 1);
 by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2);
 by (dtac ltD 1);
-by (asm_simp_tac (simpset() addsplits [split_if]) 1);
-by (Blast_tac 1);
+by Auto_tac;
 qed "mod2_cases";
 
-Goal "!!m. m:nat ==> succ(succ(m)) mod 2 = m mod 2";
+Goal "m:nat ==> succ(succ(m)) mod 2 = m mod 2";
 by (subgoal_tac "m mod 2: 2" 1);
 by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2);
 by (asm_simp_tac (simpset() addsimps [mod_succ] setloop Step_tac) 1);
 qed "mod2_succ_succ";
 
-Goal "!!m. m:nat ==> (m#+m) mod 2 = 0";
+Goal "m:nat ==> (m#+m) mod 2 = 0";
 by (etac nat_induct 1);
 by (simp_tac (simpset() addsimps [mod_less]) 1);
 by (asm_simp_tac (simpset() addsimps [mod2_succ_succ, add_succ_right]) 1);
@@ -446,12 +436,12 @@
 
 (**** Additional theorems about "le" ****)
 
-Goal "!!m n. [| m:nat;  n:nat |] ==> m le m #+ n";
+Goal "[| m:nat;  n:nat |] ==> m le m #+ n";
 by (etac nat_induct 1);
 by (ALLGOALS Asm_simp_tac);
 qed "add_le_self";
 
-Goal "!!m n. [| m:nat;  n:nat |] ==> m le n #+ m";
+Goal "[| m:nat;  n:nat |] ==> m le n #+ m";
 by (stac add_commute 1);
 by (REPEAT (ares_tac [add_le_self] 1));
 qed "add_le_self2";
@@ -459,7 +449,7 @@
 (*** Monotonicity of Addition ***)
 
 (*strict, in 1st argument; proof is by rule induction on 'less than'*)
-Goal "!!i j k. [| i<j; j:nat; k:nat |] ==> i#+k < j#+k";
+Goal "[| i<j; j:nat; k:nat |] ==> i#+k < j#+k";
 by (forward_tac [lt_nat_in_nat] 1);
 by (assume_tac 1);
 by (etac succ_lt_induct 1);
@@ -467,7 +457,7 @@
 qed "add_lt_mono1";
 
 (*strict, in both arguments*)
-Goal "!!i j k l. [| i<j; k<l; j:nat; l:nat |] ==> i#+k < j#+l";
+Goal "[| i<j; k<l; j:nat; l:nat |] ==> i#+k < j#+l";
 by (rtac (add_lt_mono1 RS lt_trans) 1);
 by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat] 1));
 by (EVERY [stac add_commute 1,
@@ -487,15 +477,13 @@
 qed "Ord_lt_mono_imp_le_mono";
 
 (*le monotonicity, 1st argument*)
-Goal
-    "!!i j k. [| i le j; j:nat; k:nat |] ==> i#+k le j#+k";
+Goal "[| i le j; j:nat; k:nat |] ==> i#+k le j#+k";
 by (res_inst_tac [("f", "%j. j#+k")] Ord_lt_mono_imp_le_mono 1);
 by (REPEAT (ares_tac [add_lt_mono1, add_type RS nat_into_Ord] 1));
 qed "add_le_mono1";
 
 (* le monotonicity, BOTH arguments*)
-Goal
-    "!!i j k. [| i le j; k le l; j:nat; l:nat |] ==> i#+k le j#+l";
+Goal "[| i le j; k le l; j:nat; l:nat |] ==> i#+k le j#+l";
 by (rtac (add_le_mono1 RS le_trans) 1);
 by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1));
 by (EVERY [stac add_commute 1,
@@ -506,15 +494,14 @@
 
 (*** Monotonicity of Multiplication ***)
 
-Goal "!!i j k. [| i le j; j:nat; k:nat |] ==> i#*k le j#*k";
+Goal "[| i le j; j:nat; k:nat |] ==> i#*k le j#*k";
 by (forward_tac [lt_nat_in_nat] 1);
 by (nat_ind_tac "k" [] 2);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_le_mono])));
 qed "mult_le_mono1";
 
 (* le monotonicity, BOTH arguments*)
-Goal
-    "!!i j k. [| i le j; k le l; j:nat; l:nat |] ==> i#*k le j#*l";
+Goal "[| i le j; k le l; j:nat; l:nat |] ==> i#*k le j#*l";
 by (rtac (mult_le_mono1 RS le_trans) 1);
 by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1));
 by (EVERY [stac mult_commute 1,
@@ -524,7 +511,7 @@
 qed "mult_le_mono";
 
 (*strict, in 1st argument; proof is by induction on k>0*)
-Goal "!!i j k. [| i<j; 0<k; j:nat; k:nat |] ==> k#*i < k#*j";
+Goal "[| i<j; 0<k; j:nat; k:nat |] ==> k#*i < k#*j";
 by (etac zero_lt_natE 1);
 by (forward_tac [lt_nat_in_nat] 2);
 by (ALLGOALS Asm_simp_tac);
@@ -532,21 +519,20 @@
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_lt_mono])));
 qed "mult_lt_mono2";
 
-Goal "!!k. [| i<j; 0<k; i:nat; j:nat; k:nat |] ==> i#*k < j#*k";
+Goal "[| i<j; 0<c; i:nat; j:nat; c:nat |] ==> i#*c < j#*c";
 by (asm_simp_tac (simpset() addsimps [mult_lt_mono2, mult_commute]) 1);
 qed "mult_lt_mono1";
 
-Goal "!!k. [| m: nat; n: nat |] ==> 0 < m#*n <-> 0<m & 0<n";
+Goal "[| m: nat; n: nat |] ==> 0 < m#*n <-> 0<m & 0<n";
 by (best_tac (claset() addEs [natE] addss (simpset())) 1);
 qed "zero_lt_mult_iff";
 
-Goal "!!k. [| m: nat; n: nat |] ==> m#*n = 1 <-> m=1 & n=1";
+Goal "[| m: nat; n: nat |] ==> m#*n = 1 <-> m=1 & n=1";
 by (best_tac (claset() addEs [natE] addss (simpset())) 1);
 qed "mult_eq_1_iff";
 
 (*Cancellation law for division*)
-Goal
-   "!!k. [| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> (k#*m) div (k#*n) = m div n";
+Goal "[| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> (k#*m) div (k#*n) = m div n";
 by (eres_inst_tac [("i","m")] complete_induct 1);
 by (excluded_middle_tac "x<n" 1);
 by (asm_simp_tac (simpset() addsimps [div_less, zero_lt_mult_iff, 
@@ -558,8 +544,7 @@
                          div_termination RS ltD]) 1);
 qed "div_cancel";
 
-Goal
-   "!!k. [| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> \
+Goal "[| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> \
 \        (k#*m) mod (k#*n) = k #* (m mod n)";
 by (eres_inst_tac [("i","m")] complete_induct 1);
 by (excluded_middle_tac "x<n" 1);
@@ -576,7 +561,7 @@
 
 val mono_lemma = (nat_into_Ord RS Ord_0_lt) RSN (2,mult_lt_mono2);
 
-Goal "!!m n. [| m = m#*n; m: nat; n: nat |] ==> n=1 | m=0";
+Goal "[| m = m#*n; m: nat; n: nat |] ==> n=1 | m=0";
 by (rtac disjCI 1);
 by (dtac sym 1);
 by (rtac Ord_linear_lt 1 THEN REPEAT_SOME (ares_tac [nat_into_Ord,nat_1I]));
--- a/src/ZF/Bool.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Bool.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -55,7 +55,7 @@
 
 
 Goal 
-    "!!b. [|  b: bool;  c: A(1);  d: A(0) |] ==> cond(b,c,d): A(b)";
+    "[|  b: bool;  c: A(1);  d: A(0) |] ==> cond(b,c,d): A(b)";
 by (bool_tac 1);
 qed "cond_type";
 
@@ -99,15 +99,15 @@
 
 (*** Laws for 'not' ***)
 
-Goal "!!a. a:bool ==> not(not(a)) = a";
+Goal "a:bool ==> not(not(a)) = a";
 by (bool_tac 1);
 qed "not_not";
 
-Goal "!!a b. a:bool ==> not(a and b) = not(a) or not(b)";
+Goal "a:bool ==> not(a and b) = not(a) or not(b)";
 by (bool_tac 1);
 qed "not_and";
 
-Goal "!!a b. a:bool ==> not(a or b) = not(a) and not(b)";
+Goal "a:bool ==> not(a or b) = not(a) and not(b)";
 by (bool_tac 1);
 qed "not_or";
 
@@ -115,44 +115,44 @@
 
 (*** Laws about 'and' ***)
 
-Goal "!!a. a: bool ==> a and a = a";
+Goal "a: bool ==> a and a = a";
 by (bool_tac 1);
 qed "and_absorb";
 
 Addsimps [and_absorb];
 
-Goal "!!a b. [| a: bool; b:bool |] ==> a and b = b and a";
+Goal "[| a: bool; b:bool |] ==> a and b = b and a";
 by (bool_tac 1);
 qed "and_commute";
 
-Goal "!!a. a: bool ==> (a and b) and c  =  a and (b and c)";
+Goal "a: bool ==> (a and b) and c  =  a and (b and c)";
 by (bool_tac 1);
 qed "and_assoc";
 
 Goal
- "!!a. [| a: bool; b:bool; c:bool |] ==> \
+ "[| a: bool; b:bool; c:bool |] ==> \
 \      (a or b) and c  =  (a and c) or (b and c)";
 by (bool_tac 1);
 qed "and_or_distrib";
 
 (** binary orion **)
 
-Goal "!!a. a: bool ==> a or a = a";
+Goal "a: bool ==> a or a = a";
 by (bool_tac 1);
 qed "or_absorb";
 
 Addsimps [or_absorb];
 
-Goal "!!a b. [| a: bool; b:bool |] ==> a or b = b or a";
+Goal "[| a: bool; b:bool |] ==> a or b = b or a";
 by (bool_tac 1);
 qed "or_commute";
 
-Goal "!!a. a: bool ==> (a or b) or c  =  a or (b or c)";
+Goal "a: bool ==> (a or b) or c  =  a or (b or c)";
 by (bool_tac 1);
 qed "or_assoc";
 
 Goal
- "!!a b c. [| a: bool; b: bool; c: bool |] ==> \
+ "[| a: bool; b: bool; c: bool |] ==> \
 \          (a and b) or c  =  (a or c) and (b or c)";
 by (bool_tac 1);
 qed "or_and_distrib";
--- a/src/ZF/Cardinal.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Cardinal.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -54,25 +54,25 @@
 
 (** Equipollence is an equivalence relation **)
 
-Goalw [eqpoll_def] "!!f A B. f: bij(A,B) ==> A eqpoll B";
+Goalw [eqpoll_def] "f: bij(A,B) ==> A eqpoll B";
 by (etac exI 1);
 qed "bij_imp_eqpoll";
 
 (*A eqpoll A*)
 bind_thm ("eqpoll_refl", id_bij RS bij_imp_eqpoll);
 
-Goalw [eqpoll_def] "!!X Y. X eqpoll Y ==> Y eqpoll X";
+Goalw [eqpoll_def] "X eqpoll Y ==> Y eqpoll X";
 by (blast_tac (claset() addIs [bij_converse_bij]) 1);
 qed "eqpoll_sym";
 
 Goalw [eqpoll_def]
-    "!!X Y. [| X eqpoll Y;  Y eqpoll Z |] ==> X eqpoll Z";
+    "[| X eqpoll Y;  Y eqpoll Z |] ==> X eqpoll Z";
 by (blast_tac (claset() addIs [comp_bij]) 1);
 qed "eqpoll_trans";
 
 (** Le-pollence is a partial ordering **)
 
-Goalw [lepoll_def] "!!X Y. X<=Y ==> X lepoll Y";
+Goalw [lepoll_def] "X<=Y ==> X lepoll Y";
 by (rtac exI 1);
 by (etac id_subset_inj 1);
 qed "subset_imp_lepoll";
@@ -82,18 +82,18 @@
 bind_thm ("le_imp_lepoll", le_imp_subset RS subset_imp_lepoll);
 
 Goalw [eqpoll_def, bij_def, lepoll_def]
-    "!!X Y. X eqpoll Y ==> X lepoll Y";
+    "X eqpoll Y ==> X lepoll Y";
 by (Blast_tac 1);
 qed "eqpoll_imp_lepoll";
 
 Goalw [lepoll_def]
-    "!!X Y. [| X lepoll Y;  Y lepoll Z |] ==> X lepoll Z";
+    "[| X lepoll Y;  Y lepoll Z |] ==> X lepoll Z";
 by (blast_tac (claset() addIs [comp_inj]) 1);
 qed "lepoll_trans";
 
 (*Asymmetry law*)
 Goalw [lepoll_def,eqpoll_def]
-    "!!X Y. [| X lepoll Y;  Y lepoll X |] ==> X eqpoll Y";
+    "[| X lepoll Y;  Y lepoll X |] ==> X eqpoll Y";
 by (REPEAT (etac exE 1));
 by (rtac schroeder_bernstein 1);
 by (REPEAT (assume_tac 1));
@@ -121,7 +121,7 @@
 qed "lepoll_0_iff";
 
 Goalw [lepoll_def] 
-    "!!A. [| A lepoll B; C lepoll D; B Int D = 0 |] ==> A Un C lepoll B Un D";
+    "[| A lepoll B; C lepoll D; B Int D = 0 |] ==> A Un C lepoll B Un D";
 by (blast_tac (claset() addIs [inj_disjoint_Un]) 1);
 qed "Un_lepoll_Un";
 
@@ -133,7 +133,7 @@
 qed "eqpoll_0_iff";
 
 Goalw [eqpoll_def] 
-    "!!A. [| A eqpoll B;  C eqpoll D;  A Int C = 0;  B Int D = 0 |] ==> \
+    "[| A eqpoll B;  C eqpoll D;  A Int C = 0;  B Int D = 0 |] ==> \
 \         A Un C eqpoll B Un D";
 by (blast_tac (claset() addIs [bij_disjoint_Un]) 1);
 qed "eqpoll_disjoint_Un";
@@ -141,12 +141,12 @@
 
 (*** lesspoll: contributions by Krzysztof Grabczewski ***)
 
-Goalw [lesspoll_def] "!!A. A lesspoll B ==> A lepoll B";
+Goalw [lesspoll_def] "A lesspoll B ==> A lepoll B";
 by (Blast_tac 1);
 qed "lesspoll_imp_lepoll";
 
 Goalw [lepoll_def]
-        "!!A. [| A lepoll B; well_ord(B,r) |] ==> EX s. well_ord(A,s)";
+        "[| A lepoll B; well_ord(B,r) |] ==> EX s. well_ord(A,s)";
 by (blast_tac (claset() addIs [well_ord_rvimage]) 1);
 qed "lepoll_well_ord";
 
@@ -155,31 +155,31 @@
 qed "lepoll_iff_leqpoll";
 
 Goalw [inj_def, surj_def] 
-  "!!f. [| f : inj(A, succ(m)); f ~: surj(A, succ(m)) |] ==> EX f. f:inj(A,m)";
+  "[| f : inj(A, succ(m)); f ~: surj(A, succ(m)) |] ==> EX f. f:inj(A,m)";
 by (safe_tac (claset_of ZF.thy));
 by (swap_res_tac [exI] 1);
 by (res_inst_tac [("a", "lam z:A. if(f`z=m, y, f`z)")] CollectI 1);
 by (best_tac (claset() addSIs [if_type RS lam_type]
                        addEs [apply_funtype RS succE]) 1);
 (*Proving it's injective*)
-by (asm_simp_tac (simpset() addsplits [split_if]) 1);
+by (Asm_simp_tac 1);
 by (blast_tac (claset() delrules [equalityI]) 1);
 qed "inj_not_surj_succ";
 
 (** Variations on transitivity **)
 
 Goalw [lesspoll_def]
-      "!!X. [| X lesspoll Y; Y lesspoll Z |] ==> X lesspoll Z";
+      "[| X lesspoll Y; Y lesspoll Z |] ==> X lesspoll Z";
 by (blast_tac (claset() addSEs [eqpollE] addIs [eqpollI, lepoll_trans]) 1);
 qed "lesspoll_trans";
 
 Goalw [lesspoll_def]
-      "!!X. [| X lesspoll Y; Y lepoll Z |] ==> X lesspoll Z";
+      "[| X lesspoll Y; Y lepoll Z |] ==> X lesspoll Z";
 by (blast_tac (claset() addSEs [eqpollE] addIs [eqpollI, lepoll_trans]) 1);
 qed "lesspoll_lepoll_lesspoll";
 
 Goalw [lesspoll_def] 
-      "!!X. [| X lesspoll Y; Z lepoll X |] ==> Z lesspoll Y";
+      "[| X lesspoll Y; Z lepoll X |] ==> Z lesspoll Y";
 by (blast_tac (claset() addSEs [eqpollE] addIs [eqpollI, lepoll_trans]) 1);
 qed "lepoll_lesspoll_lesspoll";
 
@@ -195,7 +195,7 @@
 by (ALLGOALS (blast_tac (claset() addSIs [premP] addSDs [premNot])));
 qed "Least_equality";
 
-Goal "!!i. [| P(i);  Ord(i) |] ==> P(LEAST x. P(x))";
+Goal "[| P(i);  Ord(i) |] ==> P(LEAST x. P(x))";
 by (etac rev_mp 1);
 by (trans_ind_tac "i" [] 1);
 by (rtac impI 1);
@@ -206,7 +206,7 @@
 qed "LeastI";
 
 (*Proof is almost identical to the one above!*)
-Goal "!!i. [| P(i);  Ord(i) |] ==> (LEAST x. P(x)) le i";
+Goal "[| P(i);  Ord(i) |] ==> (LEAST x. P(x)) le i";
 by (etac rev_mp 1);
 by (trans_ind_tac "i" [] 1);
 by (rtac impI 1);
@@ -217,7 +217,7 @@
 qed "Least_le";
 
 (*LEAST really is the smallest*)
-Goal "!!i. [| P(i);  i < (LEAST x. P(x)) |] ==> Q";
+Goal "[| P(i);  i < (LEAST x. P(x)) |] ==> Q";
 by (rtac (Least_le RSN (2,lt_trans2) RS lt_irrefl) 1);
 by (REPEAT (eresolve_tac [asm_rl, ltE] 1));
 qed "less_LeastE";
@@ -232,7 +232,7 @@
 
 (*If there is no such P then LEAST is vacuously 0*)
 Goalw [Least_def]
-    "!!P. [| ~ (EX i. Ord(i) & P(i)) |] ==> (LEAST x. P(x)) = 0";
+    "[| ~ (EX i. Ord(i) & P(i)) |] ==> (LEAST x. P(x)) = 0";
 by (rtac the_0 1);
 by (Blast_tac 1);
 qed "Least_0";
@@ -264,7 +264,7 @@
 
 (*Under AC, the premise becomes trivial; one consequence is ||A|| = |A|*)
 Goalw [cardinal_def]
-    "!!A. well_ord(A,r) ==> |A| eqpoll A";
+    "well_ord(A,r) ==> |A| eqpoll A";
 by (rtac LeastI 1);
 by (etac Ord_ordertype 2);
 by (etac (ordermap_bij RS bij_converse_bij RS bij_imp_eqpoll) 1);
@@ -274,25 +274,25 @@
 bind_thm ("Ord_cardinal_eqpoll", well_ord_Memrel RS well_ord_cardinal_eqpoll);
 
 Goal
-    "!!X Y. [| well_ord(X,r);  well_ord(Y,s);  |X| = |Y| |] ==> X eqpoll Y";
+    "[| well_ord(X,r);  well_ord(Y,s);  |X| = |Y| |] ==> X eqpoll Y";
 by (rtac (eqpoll_sym RS eqpoll_trans) 1);
 by (etac well_ord_cardinal_eqpoll 1);
 by (asm_simp_tac (simpset() addsimps [well_ord_cardinal_eqpoll]) 1);
 qed "well_ord_cardinal_eqE";
 
 Goal
-    "!!X Y. [| well_ord(X,r);  well_ord(Y,s) |] ==> |X| = |Y| <-> X eqpoll Y";
+    "[| well_ord(X,r);  well_ord(Y,s) |] ==> |X| = |Y| <-> X eqpoll Y";
 by (blast_tac (claset() addIs [cardinal_cong, well_ord_cardinal_eqE]) 1);
 qed "well_ord_cardinal_eqpoll_iff";
 
 
 (** Observations from Kunen, page 28 **)
 
-Goalw [cardinal_def] "!!i. Ord(i) ==> |i| le i";
+Goalw [cardinal_def] "Ord(i) ==> |i| le i";
 by (etac (eqpoll_refl RS Least_le) 1);
 qed "Ord_cardinal_le";
 
-Goalw [Card_def] "!!K. Card(K) ==> |K| = K";
+Goalw [Card_def] "Card(K) ==> |K| = K";
 by (etac sym 1);
 qed "Card_cardinal_eq";
 
@@ -308,7 +308,7 @@
 by (rtac Ord_Least 1);
 qed "Card_is_Ord";
 
-Goal "!!K. Card(K) ==> K le |K|";
+Goal "Card(K) ==> K le |K|";
 by (asm_simp_tac (simpset() addsimps [le_refl, Card_is_Ord, Card_cardinal_eq]) 1);
 qed "Card_cardinal_le";
 
@@ -326,7 +326,7 @@
 by (ALLGOALS assume_tac);
 qed "Card_iff_initial";
 
-Goalw [lesspoll_def] "!!a. [| Card(a); i<a |] ==> i lesspoll a";
+Goalw [lesspoll_def] "[| Card(a); i<a |] ==> i lesspoll a";
 by (dresolve_tac [Card_iff_initial RS iffD1] 1);
 by (blast_tac (claset() addSIs [leI RS le_imp_lepoll]) 1);
 qed "lt_Card_imp_lesspoll";
@@ -359,7 +359,7 @@
 qed "Card_cardinal";
 
 (*Kunen's Lemma 10.5*)
-Goal "!!i j. [| |i| le j;  j le i |] ==> |j| = |i|";
+Goal "[| |i| le j;  j le i |] ==> |j| = |i|";
 by (rtac (eqpollI RS cardinal_cong) 1);
 by (etac le_imp_lepoll 1);
 by (rtac lepoll_trans 1);
@@ -369,7 +369,7 @@
 by (REPEAT (eresolve_tac [ltE, Ord_succD] 1));
 qed "cardinal_eq_lemma";
 
-Goal "!!i j. i le j ==> |i| le |j|";
+Goal "i le j ==> |i| le |j|";
 by (res_inst_tac [("i","|i|"),("j","|j|")] Ord_linear_le 1);
 by (REPEAT_FIRST (ares_tac [Ord_cardinal, le_eqI]));
 by (rtac cardinal_eq_lemma 1);
@@ -380,23 +380,23 @@
 qed "cardinal_mono";
 
 (*Since we have |succ(nat)| le |nat|, the converse of cardinal_mono fails!*)
-Goal "!!i j. [| |i| < |j|;  Ord(i);  Ord(j) |] ==> i < j";
+Goal "[| |i| < |j|;  Ord(i);  Ord(j) |] ==> i < j";
 by (rtac Ord_linear2 1);
 by (REPEAT_SOME assume_tac);
 by (etac (lt_trans2 RS lt_irrefl) 1);
 by (etac cardinal_mono 1);
 qed "cardinal_lt_imp_lt";
 
-Goal "!!i j. [| |i| < K;  Ord(i);  Card(K) |] ==> i < K";
+Goal "[| |i| < K;  Ord(i);  Card(K) |] ==> i < K";
 by (asm_simp_tac (simpset() addsimps 
                   [cardinal_lt_imp_lt, Card_is_Ord, Card_cardinal_eq]) 1);
 qed "Card_lt_imp_lt";
 
-Goal "!!i j. [| Ord(i);  Card(K) |] ==> (|i| < K) <-> (i < K)";
+Goal "[| Ord(i);  Card(K) |] ==> (|i| < K) <-> (i < K)";
 by (blast_tac (claset() addIs [Card_lt_imp_lt, Ord_cardinal_le RS lt_trans1]) 1);
 qed "Card_lt_iff";
 
-Goal "!!i j. [| Ord(i);  Card(K) |] ==> (K le |i|) <-> (K le i)";
+Goal "[| Ord(i);  Card(K) |] ==> (K le |i|) <-> (K le i)";
 by (asm_simp_tac (simpset() addsimps 
                   [Card_lt_iff, Card_is_Ord, Ord_cardinal, 
                    not_lt_iff_le RS iff_sym]) 1);
@@ -404,7 +404,7 @@
 
 (*Can use AC or finiteness to discharge first premise*)
 Goal
-    "!!A B. [| well_ord(B,r);  A lepoll B |] ==> |A| le |B|";
+    "[| well_ord(B,r);  A lepoll B |] ==> |A| le |B|";
 by (res_inst_tac [("i","|A|"),("j","|B|")] Ord_linear_le 1);
 by (REPEAT_FIRST (ares_tac [Ord_cardinal, le_eqI]));
 by (rtac (eqpollI RS cardinal_cong) 1 THEN assume_tac 1);
@@ -421,7 +421,7 @@
 qed "well_ord_lepoll_imp_Card_le";
 
 
-Goal "!!A. [| A lepoll i; Ord(i) |] ==> |A| le i";
+Goal "[| A lepoll i; Ord(i) |] ==> |A| le i";
 by (rtac le_trans 1);
 by (etac (well_ord_Memrel RS well_ord_lepoll_imp_Card_le) 1);
 by (assume_tac 1);
@@ -432,7 +432,7 @@
 (*** The finite cardinals ***)
 
 Goalw [lepoll_def, inj_def]
- "!!A B. [| cons(u,A) lepoll cons(v,B);  u~:A;  v~:B |] ==> A lepoll B";
+ "[| cons(u,A) lepoll cons(v,B);  u~:A;  v~:B |] ==> A lepoll B";
 by Safe_tac;
 by (res_inst_tac [("x", "lam x:A. if(f`x=v, f`u, f`x)")] exI 1);
 by (rtac CollectI 1);
@@ -441,18 +441,18 @@
 by (blast_tac (claset() addEs [apply_funtype RS consE]) 1);
 by (blast_tac (claset() addSEs [mem_irrefl] addEs [apply_funtype RS consE]) 1);
 (*Proving it's injective*)
-by (asm_simp_tac (simpset() addsplits [split_if]) 1);
+by (Asm_simp_tac 1);
 by (Blast_tac 1);
 qed "cons_lepoll_consD";
 
 Goal
- "!!A B. [| cons(u,A) eqpoll cons(v,B);  u~:A;  v~:B |] ==> A eqpoll B";
+ "[| cons(u,A) eqpoll cons(v,B);  u~:A;  v~:B |] ==> A eqpoll B";
 by (asm_full_simp_tac (simpset() addsimps [eqpoll_iff]) 1);
 by (blast_tac (claset() addIs [cons_lepoll_consD]) 1);
 qed "cons_eqpoll_consD";
 
 (*Lemma suggested by Mike Fourman*)
-Goalw [succ_def] "!!m n. succ(m) lepoll succ(n) ==> m lepoll n";
+Goalw [succ_def] "succ(m) lepoll succ(n) ==> m lepoll n";
 by (etac cons_lepoll_consD 1);
 by (REPEAT (rtac mem_not_refl 1));
 qed "succ_lepoll_succD";
@@ -471,7 +471,7 @@
 bind_thm ("nat_lepoll_imp_le", nat_lepoll_imp_le_lemma RS bspec RS mp);
 
 Goal
-    "!!m n. [| m:nat; n: nat |] ==> m eqpoll n <-> m = n";
+    "[| m:nat; n: nat |] ==> m eqpoll n <-> m = n";
 by (rtac iffI 1);
 by (asm_simp_tac (simpset() addsimps [eqpoll_refl]) 2);
 by (blast_tac (claset() addIs [nat_lepoll_imp_le, le_anti_sym] 
@@ -480,7 +480,7 @@
 
 (*The object of all this work: every natural number is a (finite) cardinal*)
 Goalw [Card_def,cardinal_def]
-    "!!n. n: nat ==> Card(n)";
+    "n: nat ==> Card(n)";
 by (stac Least_equality 1);
 by (REPEAT_FIRST (ares_tac [eqpoll_refl, nat_into_Ord, refl]));
 by (asm_simp_tac (simpset() addsimps [lt_nat_in_nat RS nat_eqpoll_iff]) 1);
@@ -488,7 +488,7 @@
 qed "nat_into_Card";
 
 (*Part of Kunen's Lemma 10.6*)
-Goal "!!n. [| succ(n) lepoll n;  n:nat |] ==> P";
+Goal "[| succ(n) lepoll n;  n:nat |] ==> P";
 by (rtac (nat_lepoll_imp_le RS lt_irrefl) 1);
 by (REPEAT (ares_tac [nat_succI] 1));
 qed "succ_lepoll_natE";
@@ -497,7 +497,7 @@
 (** lepoll, lesspoll and natural numbers **)
 
 Goalw [lesspoll_def]
-      "!!m. [| A lepoll m; m:nat |] ==> A lesspoll succ(m)";
+      "[| A lepoll m; m:nat |] ==> A lesspoll succ(m)";
 by (rtac conjI 1);
 by (blast_tac (claset() addIs [subset_imp_lepoll RSN (2,lepoll_trans)]) 1);
 by (rtac notI 1);
@@ -507,17 +507,17 @@
 qed "lepoll_imp_lesspoll_succ";
 
 Goalw [lesspoll_def, lepoll_def, eqpoll_def, bij_def]
-      "!!m. [| A lesspoll succ(m); m:nat |] ==> A lepoll m";
+      "[| A lesspoll succ(m); m:nat |] ==> A lepoll m";
 by (Clarify_tac 1);
 by (blast_tac (claset() addSIs [inj_not_surj_succ]) 1);
 qed "lesspoll_succ_imp_lepoll";
 
-Goal "!!m. m:nat ==> A lesspoll succ(m) <-> A lepoll m";
+Goal "m:nat ==> A lesspoll succ(m) <-> A lepoll m";
 by (blast_tac (claset() addSIs [lepoll_imp_lesspoll_succ, 
                             lesspoll_succ_imp_lepoll]) 1);
 qed "lesspoll_succ_iff";
 
-Goal "!!A m. [| A lepoll succ(m);  m:nat |] ==>  \
+Goal "[| A lepoll succ(m);  m:nat |] ==>  \
 \                         A lepoll m | A eqpoll succ(m)";
 by (rtac disjCI 1);
 by (rtac lesspoll_succ_imp_lepoll 1);
@@ -529,7 +529,7 @@
 (*** The first infinite cardinal: Omega, or nat ***)
 
 (*This implies Kunen's Lemma 10.6*)
-Goal "!!n. [| n<i;  n:nat |] ==> ~ i lepoll n";
+Goal "[| n<i;  n:nat |] ==> ~ i lepoll n";
 by (rtac notI 1);
 by (rtac succ_lepoll_natE 1 THEN assume_tac 2);
 by (rtac lepoll_trans 1 THEN assume_tac 2);
@@ -537,7 +537,7 @@
 by (REPEAT (ares_tac [Ord_succ_subsetI RS subset_imp_lepoll] 1));
 qed "lt_not_lepoll";
 
-Goal "!!i n. [| Ord(i);  n:nat |] ==> i eqpoll n <-> i=n";
+Goal "[| Ord(i);  n:nat |] ==> i eqpoll n <-> i=n";
 by (rtac iffI 1);
 by (asm_simp_tac (simpset() addsimps [eqpoll_refl]) 2);
 by (rtac Ord_linear_lt 1);
@@ -556,7 +556,7 @@
 qed "Card_nat";
 
 (*Allows showing that |i| is a limit cardinal*)
-Goal  "!!i. nat le i ==> nat le |i|";
+Goal  "nat le i ==> nat le |i|";
 by (rtac (Card_nat RS Card_cardinal_eq RS subst) 1);
 by (etac cardinal_mono 1);
 qed "nat_le_cardinal";
@@ -567,7 +567,7 @@
 
 (*Congruence law for  cons  under equipollence*)
 Goalw [lepoll_def]
-    "!!A B. [| A lepoll B;  b ~: B |] ==> cons(a,A) lepoll cons(b,B)";
+    "[| A lepoll B;  b ~: B |] ==> cons(a,A) lepoll cons(b,B)";
 by Safe_tac;
 by (res_inst_tac [("x", "lam y: cons(a,A).if(y=a, b, f`y)")] exI 1);
 by (res_inst_tac [("d","%z. if(z:B, converse(f)`z, a)")] 
@@ -579,18 +579,18 @@
 qed "cons_lepoll_cong";
 
 Goal
-    "!!A B. [| A eqpoll B;  a ~: A;  b ~: B |] ==> cons(a,A) eqpoll cons(b,B)";
+    "[| A eqpoll B;  a ~: A;  b ~: B |] ==> cons(a,A) eqpoll cons(b,B)";
 by (asm_full_simp_tac (simpset() addsimps [eqpoll_iff, cons_lepoll_cong]) 1);
 qed "cons_eqpoll_cong";
 
 Goal
-    "!!A B. [| a ~: A;  b ~: B |] ==> \
+    "[| a ~: A;  b ~: B |] ==> \
 \           cons(a,A) lepoll cons(b,B)  <->  A lepoll B";
 by (blast_tac (claset() addIs [cons_lepoll_cong, cons_lepoll_consD]) 1);
 qed "cons_lepoll_cons_iff";
 
 Goal
-    "!!A B. [| a ~: A;  b ~: B |] ==> \
+    "[| a ~: A;  b ~: B |] ==> \
 \           cons(a,A) eqpoll cons(b,B)  <->  A eqpoll B";
 by (blast_tac (claset() addIs [cons_eqpoll_cong, cons_eqpoll_consD]) 1);
 qed "cons_eqpoll_cons_iff";
@@ -606,37 +606,35 @@
 
 (*Congruence law for  succ  under equipollence*)
 Goalw [succ_def]
-    "!!A B. A eqpoll B ==> succ(A) eqpoll succ(B)";
+    "A eqpoll B ==> succ(A) eqpoll succ(B)";
 by (REPEAT (ares_tac [cons_eqpoll_cong, mem_not_refl] 1));
 qed "succ_eqpoll_cong";
 
 (*Congruence law for + under equipollence*)
 Goalw [eqpoll_def]
-    "!!A B C D. [| A eqpoll C;  B eqpoll D |] ==> A+B eqpoll C+D";
+    "[| A eqpoll C;  B eqpoll D |] ==> A+B eqpoll C+D";
 by (blast_tac (claset() addSIs [sum_bij]) 1);
 qed "sum_eqpoll_cong";
 
 (*Congruence law for * under equipollence*)
 Goalw [eqpoll_def]
-    "!!A B C D. [| A eqpoll C;  B eqpoll D |] ==> A*B eqpoll C*D";
+    "[| A eqpoll C;  B eqpoll D |] ==> A*B eqpoll C*D";
 by (blast_tac (claset() addSIs [prod_bij]) 1);
 qed "prod_eqpoll_cong";
 
 Goalw [eqpoll_def]
-    "!!f. [| f: inj(A,B);  A Int B = 0 |] ==> A Un (B - range(f)) eqpoll B";
+    "[| f: inj(A,B);  A Int B = 0 |] ==> A Un (B - range(f)) eqpoll B";
 by (rtac exI 1);
 by (res_inst_tac [("c", "%x. if(x:A, f`x, x)"),
                   ("d", "%y. if(y: range(f), converse(f)`y, y)")] 
     lam_bijective 1);
 by (blast_tac (claset() addSIs [if_type, inj_is_fun RS apply_type]) 1);
 by (asm_simp_tac 
-    (simpset() addsimps [inj_converse_fun RS apply_funtype]
-           addsplits [split_if]) 1);
+    (simpset() addsimps [inj_converse_fun RS apply_funtype]) 1);
 by (asm_simp_tac (simpset() addsimps [inj_is_fun RS apply_rangeI, left_inverse]
                         setloop etac UnE') 1);
 by (asm_simp_tac 
-    (simpset() addsimps [inj_converse_fun RS apply_funtype, right_inverse]
-           addsplits [split_if]) 1);
+    (simpset() addsimps [inj_converse_fun RS apply_funtype, right_inverse]) 1);
 by (blast_tac (claset() addEs [equals0D]) 1);
 qed "inj_disjoint_eqpoll";
 
@@ -646,7 +644,7 @@
 
 (*If A has at most n+1 elements and a:A then A-{a} has at most n.*)
 Goalw [succ_def]
-      "!!A a n. [| a:A;  A lepoll succ(n) |] ==> A - {a} lepoll n";
+      "[| a:A;  A lepoll succ(n) |] ==> A - {a} lepoll n";
 by (rtac cons_lepoll_consD 1);
 by (rtac mem_not_refl 3);
 by (eresolve_tac [cons_Diff RS ssubst] 1);
@@ -655,19 +653,19 @@
 
 (*If A has at least n+1 elements then A-{a} has at least n.*)
 Goalw [succ_def]
-      "!!A a n. [| succ(n) lepoll A |] ==> n lepoll A - {a}";
+      "[| succ(n) lepoll A |] ==> n lepoll A - {a}";
 by (rtac cons_lepoll_consD 1);
 by (rtac mem_not_refl 2);
 by (Blast_tac 2);
 by (blast_tac (claset() addIs [subset_imp_lepoll RSN (2, lepoll_trans)]) 1);
 qed "lepoll_Diff_sing";
 
-Goal "!!A a n. [| a:A; A eqpoll succ(n) |] ==> A - {a} eqpoll n";
+Goal "[| a:A; A eqpoll succ(n) |] ==> A - {a} eqpoll n";
 by (blast_tac (claset() addSIs [eqpollI] addSEs [eqpollE] 
                     addIs [Diff_sing_lepoll,lepoll_Diff_sing]) 1);
 qed "Diff_sing_eqpoll";
 
-Goal "!!A. [| A lepoll 1; a:A |] ==> A = {a}";
+Goal "[| A lepoll 1; a:A |] ==> A = {a}";
 by (forward_tac [Diff_sing_lepoll] 1);
 by (assume_tac 1);
 by (dtac lepoll_0_is_0 1);
@@ -679,8 +677,7 @@
 by (res_inst_tac [("d","%z. snd(z)")] lam_injective 1);
 by (split_tac [split_if] 1);
 by (blast_tac (claset() addSIs [InlI, InrI]) 1);
-by (asm_full_simp_tac (simpset() addsimps [Inl_def, Inr_def]
-                       addsplits [split_if]) 1);
+by (asm_full_simp_tac (simpset() addsimps [Inl_def, Inr_def]) 1);
 qed "Un_lepoll_sum";
 
 
@@ -691,15 +688,15 @@
 qed "Finite_0";
 
 Goalw [Finite_def]
-    "!!A. [| A lepoll n;  n:nat |] ==> Finite(A)";
+    "[| A lepoll n;  n:nat |] ==> Finite(A)";
 by (etac rev_mp 1);
 by (etac nat_induct 1);
 by (blast_tac (claset() addSDs [lepoll_0_is_0] addSIs [eqpoll_refl,nat_0I]) 1);
-by (blast_tac (claset() addSDs [lepoll_succ_disj] addSIs [nat_succI]) 1);
+by (blast_tac (claset() addSDs [lepoll_succ_disj]) 1);
 qed "lepoll_nat_imp_Finite";
 
 Goalw [Finite_def]
-     "!!X. [| Y lepoll X;  Finite(X) |] ==> Finite(Y)";
+     "[| Y lepoll X;  Finite(X) |] ==> Finite(Y)";
 by (blast_tac 
     (claset() addSEs [eqpollE] 
              addIs [lepoll_trans RS 
@@ -710,7 +707,7 @@
 
 bind_thm ("Finite_Diff", Diff_subset RS subset_Finite);
 
-Goalw [Finite_def] "!!x. Finite(x) ==> Finite(cons(y,x))";
+Goalw [Finite_def] "Finite(x) ==> Finite(cons(y,x))";
 by (excluded_middle_tac "y:x" 1);
 by (asm_simp_tac (simpset() addsimps [cons_absorb]) 2);
 by (etac bexE 1);
@@ -720,19 +717,19 @@
     (simpset() addsimps [succ_def, cons_eqpoll_cong, mem_not_refl]) 1);
 qed "Finite_cons";
 
-Goalw [succ_def] "!!x. Finite(x) ==> Finite(succ(x))";
+Goalw [succ_def] "Finite(x) ==> Finite(succ(x))";
 by (etac Finite_cons 1);
 qed "Finite_succ";
 
 Goalw [Finite_def] 
-      "!!i. [| Ord(i);  ~ Finite(i) |] ==> nat le i";
+      "[| Ord(i);  ~ Finite(i) |] ==> nat le i";
 by (eresolve_tac [Ord_nat RSN (2,Ord_linear2)] 1);
 by (assume_tac 2);
 by (blast_tac (claset() addSIs [eqpoll_refl] addSEs [ltE]) 1);
 qed "nat_le_infinite_Ord";
 
 Goalw [Finite_def, eqpoll_def]
-    "!!A. Finite(A) ==> EX r. well_ord(A,r)";
+    "Finite(A) ==> EX r. well_ord(A,r)";
 by (blast_tac (claset() addIs [well_ord_rvimage, bij_is_inj, well_ord_Memrel, 
 			      nat_into_Ord]) 1);
 qed "Finite_imp_well_ord";
@@ -753,7 +750,7 @@
 by (Blast.depth_tac (claset()) 4 1);
 qed "nat_wf_on_converse_Memrel";
 
-Goal "!!n. n:nat ==> well_ord(n,converse(Memrel(n)))";
+Goal "n:nat ==> well_ord(n,converse(Memrel(n)))";
 by (forward_tac [transfer thy Ord_nat RS Ord_in_Ord RS well_ord_Memrel] 1);
 by (rewtac well_ord_def);
 by (blast_tac (claset() addSIs [tot_ord_converse, 
@@ -761,7 +758,7 @@
 qed "nat_well_ord_converse_Memrel";
 
 Goal
-    "!!A. [| well_ord(A,r);     \
+    "[| well_ord(A,r);     \
 \            well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r)))) \
 \         |] ==> well_ord(A,converse(r))";
 by (resolve_tac [well_ord_Int_iff RS iffD1] 1);
@@ -773,7 +770,7 @@
 qed "well_ord_converse";
 
 Goal
-    "!!A. [| well_ord(A,r);  A eqpoll n;  n:nat |] ==> ordertype(A,r)=n";
+    "[| well_ord(A,r);  A eqpoll n;  n:nat |] ==> ordertype(A,r)=n";
 by (rtac (Ord_ordertype RS Ord_nat_eqpoll_iff RS iffD1) 1 THEN 
     REPEAT (assume_tac 1));
 by (rtac eqpoll_trans 1 THEN assume_tac 2);
@@ -782,7 +779,7 @@
 qed "ordertype_eq_n";
 
 Goalw [Finite_def]
-    "!!A. [| Finite(A);  well_ord(A,r) |] ==> well_ord(A,converse(r))";
+    "[| Finite(A);  well_ord(A,r) |] ==> well_ord(A,converse(r))";
 by (rtac well_ord_converse 1 THEN assume_tac 1);
 by (blast_tac (claset() addDs [ordertype_eq_n] 
                        addSIs [nat_well_ord_converse_Memrel]) 1);
--- a/src/ZF/CardinalArith.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/CardinalArith.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -56,7 +56,7 @@
 by (rtac bij_0_sum 1);
 qed "sum_0_eqpoll";
 
-Goalw [cadd_def] "!!K. Card(K) ==> 0 |+| K = K";
+Goalw [cadd_def] "Card(K) ==> 0 |+| K = K";
 by (asm_simp_tac (simpset() addsimps [sum_0_eqpoll RS cardinal_cong, 
 				      Card_cardinal_eq]) 1);
 qed "cadd_0";
@@ -220,7 +220,7 @@
 qed "prod_square_lepoll";
 
 (*Could probably weaken the premise to well_ord(K,r), or remove using AC*)
-Goalw [cmult_def] "!!K. Card(K) ==> K le K |*| K";
+Goalw [cmult_def] "Card(K) ==> K le K |*| K";
 by (rtac le_trans 1);
 by (rtac well_ord_lepoll_imp_Card_le 2);
 by (rtac prod_square_lepoll 3);
@@ -295,7 +295,7 @@
 				      nat_cadd_eq_add]) 1);
 qed "nat_cmult_eq_mult";
 
-Goal "!!m n. Card(n) ==> 2 |*| n = n |+| n";
+Goal "Card(n) ==> 2 |*| n = n |+| n";
 by (asm_simp_tac 
     (simpset() addsimps [Ord_0, Ord_succ, cmult_0, cmult_succ_lemma, 
 			 Card_is_Ord, cadd_0,
@@ -318,24 +318,23 @@
 by (res_inst_tac [("d", "%y. if(y: range(f),    \
 \                               nat_case(u, %z. f`z, converse(f)`y), y)")] 
     lam_injective 1);
-by (fast_tac (claset() addSIs [if_type, nat_succI, apply_type]
-                      addIs  [inj_is_fun, inj_converse_fun]) 1);
+by (fast_tac (claset() addSIs [if_type, apply_type]
+                       addIs  [inj_is_fun, inj_converse_fun]) 1);
 by (asm_simp_tac 
     (simpset() addsimps [inj_is_fun RS apply_rangeI,
 			 inj_converse_fun RS apply_rangeI,
 			 inj_converse_fun RS apply_funtype,
 			 left_inverse, right_inverse, nat_0I, nat_succI, 
-			 nat_case_0, nat_case_succ]
-               addsplits [split_if]) 1);
+			 nat_case_0, nat_case_succ]) 1);
 qed "nat_cons_lepoll";
 
-Goal "!!i. nat lepoll A ==> cons(u,A) eqpoll A";
+Goal "nat lepoll A ==> cons(u,A) eqpoll A";
 by (etac (nat_cons_lepoll RS eqpollI) 1);
 by (rtac (subset_consI RS subset_imp_lepoll) 1);
 qed "nat_cons_eqpoll";
 
 (*Specialized version required below*)
-Goalw [succ_def] "!!i. nat <= A ==> succ(A) eqpoll A";
+Goalw [succ_def] "nat <= A ==> succ(A) eqpoll A";
 by (eresolve_tac [subset_imp_lepoll RS nat_cons_eqpoll] 1);
 qed "nat_succ_eqpoll";
 
@@ -343,7 +342,7 @@
 by (blast_tac (claset() addIs [Card_nat, le_refl, Card_is_Ord]) 1);
 qed "InfCard_nat";
 
-Goalw [InfCard_def] "!!K. InfCard(K) ==> Card(K)";
+Goalw [InfCard_def] "InfCard(K) ==> Card(K)";
 by (etac conjunct1 1);
 qed "InfCard_is_Card";
 
@@ -354,7 +353,7 @@
 qed "InfCard_Un";
 
 (*Kunen's Lemma 10.11*)
-Goalw [InfCard_def] "!!K. InfCard(K) ==> Limit(K)";
+Goalw [InfCard_def] "InfCard(K) ==> Limit(K)";
 by (etac conjE 1);
 by (forward_tac [Card_is_Ord] 1);
 by (rtac (ltI RS non_succ_LimitI) 1);
@@ -509,7 +508,7 @@
 qed "ordertype_csquare_le";
 
 (*Main result: Kunen's Theorem 10.12*)
-Goal "!!K. InfCard(K) ==> K |*| K = K";
+Goal "InfCard(K) ==> K |*| K = K";
 by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1);
 by (etac rev_mp 1);
 by (trans_ind_tac "K" [] 1);
@@ -539,7 +538,7 @@
 
 (** Toward's Kunen's Corollary 10.13 (1) **)
 
-Goal "!!K. [| InfCard(K);  L le K;  0<L |] ==> K |*| L = K";
+Goal "[| InfCard(K);  L le K;  0<L |] ==> K |*| L = K";
 by (rtac le_anti_sym 1);
 by (etac ltE 2 THEN
     REPEAT (ares_tac [cmult_le_self, InfCard_is_Card] 2));
@@ -562,7 +561,7 @@
 qed "InfCard_cmult_eq";
 
 (*This proof appear to be the simplest!*)
-Goal "!!K. InfCard(K) ==> K |+| K = K";
+Goal "InfCard(K) ==> K |+| K = K";
 by (asm_simp_tac
     (simpset() addsimps [cmult_2 RS sym, InfCard_is_Card, cmult_commute]) 1);
 by (rtac InfCard_le_cmult_eq 1);
@@ -571,7 +570,7 @@
 qed "InfCard_cdouble_eq";
 
 (*Corollary 10.13 (1), for cardinal addition*)
-Goal "!!K. [| InfCard(K);  L le K |] ==> K |+| L = K";
+Goal "[| InfCard(K);  L le K |] ==> K |+| L = K";
 by (rtac le_anti_sym 1);
 by (etac ltE 2 THEN
     REPEAT (ares_tac [cadd_le_self, InfCard_is_Card] 2));
@@ -620,7 +619,7 @@
 qed "jump_cardinal_iff";
 
 (*The easy part of Theorem 10.16: jump_cardinal(K) exceeds K*)
-Goal "!!K. Ord(K) ==> K < jump_cardinal(K)";
+Goal "Ord(K) ==> K < jump_cardinal(K)";
 by (resolve_tac [Ord_jump_cardinal RSN (2,ltI)] 1);
 by (resolve_tac [jump_cardinal_iff RS iffD2] 1);
 by (REPEAT_FIRST (ares_tac [exI, conjI, well_ord_Memrel]));
@@ -669,7 +668,7 @@
 
 bind_thm ("lt_csucc", csucc_basic RS conjunct2);
 
-Goal "!!K. Ord(K) ==> 0 < csucc(K)";
+Goal "Ord(K) ==> 0 < csucc(K)";
 by (resolve_tac [[Ord_0_le, lt_csucc] MRS lt_trans1] 1);
 by (REPEAT (assume_tac 1));
 qed "Ord_0_lt_csucc";
@@ -729,11 +728,11 @@
 by (blast_tac (claset() addIs [bij_converse_bij RS bij_is_fun RS apply_type]) 1);
 val lemma = result();
 
-Goalw [Finite_def] "!!A. Finite(A) ==> A : Fin(A)";
+Goalw [Finite_def] "Finite(A) ==> A : Fin(A)";
 by (blast_tac (claset() addIs [lemma RS spec RS mp]) 1);
 qed "Finite_into_Fin";
 
-Goal "!!A. A : Fin(U) ==> Finite(A)";
+Goal "A : Fin(U) ==> Finite(A)";
 by (fast_tac (claset() addSIs [Finite_0, Finite_cons] addEs [Fin.induct]) 1);
 qed "Fin_into_Finite";
 
@@ -783,7 +782,7 @@
 qed "Finite_imp_cardinal_cons";
 
 
-Goal "!!a A. [| Finite(A);  a:A |] ==> succ(|A-{a}|) = |A|";
+Goal "[| Finite(A);  a:A |] ==> succ(|A-{a}|) = |A|";
 by (res_inst_tac [("b", "A")] (cons_Diff RS subst) 1);
 by (assume_tac 1);
 by (asm_simp_tac (simpset() addsimps [Finite_imp_cardinal_cons,
@@ -791,7 +790,7 @@
 by (asm_simp_tac (simpset() addsimps [cons_Diff]) 1);
 qed "Finite_imp_succ_cardinal_Diff";
 
-Goal "!!a A. [| Finite(A);  a:A |] ==> |A-{a}| < |A|";
+Goal "[| Finite(A);  a:A |] ==> |A-{a}| < |A|";
 by (rtac succ_leE 1);
 by (asm_simp_tac (simpset() addsimps [Finite_imp_succ_cardinal_Diff, 
 				      Ord_cardinal RS le_refl]) 1);
@@ -803,7 +802,7 @@
 val nat_implies_well_ord =
   (transfer CardinalArith.thy nat_into_Ord) RS well_ord_Memrel;
 
-Goal "!!m n. [| m:nat; n:nat |] ==> m + n eqpoll m #+ n";
+Goal "[| m:nat; n:nat |] ==> m + n eqpoll m #+ n";
 by (rtac eqpoll_trans 1);
 by (resolve_tac [well_ord_radd RS well_ord_cardinal_eqpoll RS eqpoll_sym] 1);
 by (REPEAT (etac nat_implies_well_ord 1));
@@ -812,6 +811,6 @@
 qed "nat_sum_eqpoll_sum";
 
 goal Nat.thy "!!m. [| m le n; n:nat |] ==> m:nat";
-by (blast_tac (claset() addSIs [nat_succI] addSDs [lt_nat_in_nat]) 1);
+by (blast_tac (claset() addSDs [lt_nat_in_nat]) 1);
 qed "le_in_nat";
 
--- a/src/ZF/Cardinal_AC.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Cardinal_AC.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -19,7 +19,7 @@
 
 val cardinal_idem = cardinal_eqpoll RS cardinal_cong;
 
-Goal "!!X Y. |X| = |Y| ==> X eqpoll Y";
+Goal "|X| = |Y| ==> X eqpoll Y";
 by (resolve_tac [AC_well_ord RS exE] 1);
 by (resolve_tac [AC_well_ord RS exE] 1);
 by (rtac well_ord_cardinal_eqE 1);
@@ -37,7 +37,7 @@
                                        eqpoll_disjoint_Un]) 1);
 qed "cardinal_disjoint_Un";
 
-Goal "!!A B. A lepoll B ==> |A| le |B|";
+Goal "A lepoll B ==> |A| le |B|";
 by (resolve_tac [AC_well_ord RS exE] 1);
 by (etac well_ord_lepoll_imp_Card_le 1);
 by (assume_tac 1);
@@ -67,7 +67,7 @@
 by (REPEAT_SOME assume_tac);
 qed "cadd_cmult_distrib";
 
-Goal "!!A. InfCard(|A|) ==> A*A eqpoll A";
+Goal "InfCard(|A|) ==> A*A eqpoll A";
 by (resolve_tac [AC_well_ord RS exE] 1);
 by (etac well_ord_InfCard_square_eq 1);
 by (assume_tac 1);
@@ -76,20 +76,20 @@
 
 (*** Other applications of AC ***)
 
-Goal "!!A B. |A| le |B| ==> A lepoll B";
+Goal "|A| le |B| ==> A lepoll B";
 by (resolve_tac [cardinal_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll RS 
                  lepoll_trans] 1);
 by (eresolve_tac [le_imp_subset RS subset_imp_lepoll RS lepoll_trans] 1);
 by (resolve_tac [cardinal_eqpoll RS eqpoll_imp_lepoll] 1);
 qed "Card_le_imp_lepoll";
 
-Goal "!!A K. Card(K) ==> |A| le K <-> A lepoll K";
+Goal "Card(K) ==> |A| le K <-> A lepoll K";
 by (eresolve_tac [Card_cardinal_eq RS subst] 1 THEN
     rtac iffI 1 THEN
     DEPTH_SOLVE (eresolve_tac [Card_le_imp_lepoll,lepoll_imp_Card_le] 1));
 qed "le_Card_iff";
 
-Goalw [surj_def] "!!f. f: surj(X,Y) ==> EX g. g: inj(Y,X)";
+Goalw [surj_def] "f: surj(X,Y) ==> EX g. g: inj(Y,X)";
 by (etac CollectE 1);
 by (res_inst_tac [("A1", "Y"), ("B1", "%y. f-``{y}")] (AC_Pi RS exE) 1);
 by (fast_tac (claset() addSEs [apply_Pair]) 1);
@@ -98,7 +98,7 @@
 qed "surj_implies_inj";
 
 (*Kunen's Lemma 10.20*)
-Goal "!!f. f: surj(X,Y) ==> |Y| le |X|";
+Goal "f: surj(X,Y) ==> |Y| le |X|";
 by (rtac lepoll_imp_Card_le 1);
 by (eresolve_tac [surj_implies_inj RS exE] 1);
 by (rewtac lepoll_def);
@@ -192,7 +192,6 @@
                   MRS lt_subset_trans] 1);
 by (REPEAT (assume_tac 1));
 by (blast_tac (claset() addSIs [Ord_UN] addEs [ltE]) 2);
-by (asm_simp_tac (simpset() addsimps [inj_converse_fun RS apply_type]
-                        addsplits [split_if]) 1);
+by (asm_simp_tac (simpset() addsimps [inj_converse_fun RS apply_type]) 1);
 qed "le_UN_Ord_lt_csucc";
 
--- a/src/ZF/Coind/Map.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Coind/Map.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -36,15 +36,15 @@
 (* Lemmas                                                       *)
 (* ############################################################ *)
 
-Goal "!!A. a:A ==> Sigma(A,B)``{a} = B(a)";
+Goal "a:A ==> Sigma(A,B)``{a} = B(a)";
 by (Fast_tac 1);
 qed "qbeta";
 
-Goal "!!A. a~:A ==> Sigma(A,B)``{a} = 0";
+Goal "a~:A ==> Sigma(A,B)``{a} = 0";
 by (Fast_tac 1);
 qed "qbeta_emp";
 
-Goal "!!A. a ~: A ==> Sigma(A,B)``{a}=0";
+Goal "a ~: A ==> Sigma(A,B)``{a}=0";
 by (Fast_tac 1);
 qed "image_Sigma1";
 
@@ -113,8 +113,7 @@
 by (rtac (excluded_middle RS disjE) 1);
 by (etac image_Sigma1 1);
 by (dres_inst_tac [("psi", "?uu ~: B")] asm_rl 1);
-by (asm_full_simp_tac
-    (simpset() addsimps [qbeta] addsplits [split_if]) 1);
+by (asm_full_simp_tac (simpset() addsimps [qbeta]) 1);
 by Safe_tac;
 by (dres_inst_tac [("psi", "?uu ~: B")] asm_rl 3);
 by (ALLGOALS Asm_full_simp_tac);
--- a/src/ZF/Coind/Values.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Coind/Values.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -75,7 +75,7 @@
 
 (* Proving that the empty set is not a value *)
 
-Goal "!!v. v:Val ==> v ~= 0";
+Goal "v:Val ==> v ~= 0";
 by (etac ValE 1);
 by (ALLGOALS hyp_subst_tac);
 by (etac v_constNE 1);
--- a/src/ZF/Epsilon.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Epsilon.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -89,7 +89,7 @@
 by (blast_tac (claset() addIs [step,ecloseD]) 1);
 qed "eclose_induct_down";
 
-Goal "!!X. Transset(X) ==> eclose(X) = X";
+Goal "Transset(X) ==> eclose(X) = X";
 by (etac ([eclose_least, arg_subset_eclose] MRS equalityI) 1);
 by (rtac subset_refl 1);
 qed "Transset_eclose_eq_arg";
@@ -98,7 +98,7 @@
 (*** Epsilon recursion ***)
 
 (*Unused...*)
-Goal "!!A B C. [| A: eclose(B);  B: eclose(C) |] ==> A: eclose(C)";
+Goal "[| A: eclose(B);  B: eclose(C) |] ==> A: eclose(C)";
 by (rtac ([Transset_eclose, eclose_subset] MRS eclose_least RS subsetD) 1);
 by (REPEAT (assume_tac 1));
 qed "mem_eclose_trans";
@@ -158,7 +158,7 @@
 by (REPEAT (ares_tac (prems @ [lam_type]) 1 ORELSE etac bspec 1));
 qed "transrec_type";
 
-Goal "!!i. Ord(i) ==> eclose({i}) <= succ(i)";
+Goal "Ord(i) ==> eclose({i}) <= succ(i)";
 by (etac (Ord_is_Transset RS Transset_succ RS eclose_least) 1);
 by (rtac (succI1 RS singleton_subsetI) 1);
 qed "eclose_sing_Ord";
@@ -195,7 +195,7 @@
 by (asm_simp_tac (simpset() addsimps [Ord_equality]) 1);
 qed "rank_of_Ord";
 
-Goal "!!a b. a:b ==> rank(a) < rank(b)";
+Goal "a:b ==> rank(a) < rank(b)";
 by (res_inst_tac [("a1","b")] (rank RS ssubst) 1);
 by (etac (UN_I RS ltI) 1);
 by (rtac succI1 1);
@@ -209,7 +209,7 @@
 by (assume_tac 1);
 qed "eclose_rank_lt";
 
-Goal "!!a b. a<=b ==> rank(a) le rank(b)";
+Goal "a<=b ==> rank(a) le rank(b)";
 by (rtac subset_imp_le 1);
 by (stac rank 1);
 by (stac rank 1);
@@ -271,12 +271,12 @@
 
 (*** Corollaries of leastness ***)
 
-Goal "!!A B. A:B ==> eclose(A)<=eclose(B)";
+Goal "A:B ==> eclose(A)<=eclose(B)";
 by (rtac (Transset_eclose RS eclose_least) 1);
 by (etac (arg_into_eclose RS eclose_subset) 1);
 qed "mem_eclose_subset";
 
-Goal "!!A B. A<=B ==> eclose(A) <= eclose(B)";
+Goal "A<=B ==> eclose(A) <= eclose(B)";
 by (rtac (Transset_eclose RS eclose_least) 1);
 by (etac subset_trans 1);
 by (rtac arg_subset_eclose 1);
@@ -304,14 +304,13 @@
 
 Goal "transrec2(succ(i),a,b) = b(i, transrec2(i,a,b))";
 by (rtac (transrec2_def RS def_transrec RS trans) 1);
-by (simp_tac (simpset() addsimps [succ_not_0, THE_eq, if_P]
-                    addsplits [split_if]) 1);
+by (simp_tac (simpset() addsimps [succ_not_0, THE_eq, if_P]) 1);
 by (Blast_tac 1);
 qed "transrec2_succ";
 
-Goal "!!i. Limit(i) ==> transrec2(i,a,b) = (UN j<i. transrec2(j,a,b))";
+Goal "Limit(i) ==> transrec2(i,a,b) = (UN j<i. transrec2(j,a,b))";
 by (rtac (transrec2_def RS def_transrec RS trans) 1);
-by (simp_tac (simpset() addsplits [split_if]) 1);
+by (Simp_tac 1);
 by (blast_tac (claset() addSDs [Limit_has_0] addSEs [succ_LimitE]) 1);
 qed "transrec2_Limit";
 
--- a/src/ZF/EquivClass.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/EquivClass.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -74,7 +74,7 @@
 by (Blast_tac 1);
 qed "equiv_class_nondisjoint";
 
-Goalw [equiv_def] "!!A r. equiv(A,r) ==> r <= A*A";
+Goalw [equiv_def] "equiv(A,r) ==> r <= A*A";
 by (safe_tac subset_cs);
 qed "equiv_type";
 
--- a/src/ZF/Finite.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Finite.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -14,7 +14,7 @@
 
 (*** Finite powerset operator ***)
 
-Goalw Fin.defs "!!A B. A<=B ==> Fin(A) <= Fin(B)";
+Goalw Fin.defs "A<=B ==> Fin(A) <= Fin(B)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac Fin.bnd_mono 1));
 by (REPEAT (ares_tac (Pow_mono::basic_monos) 1));
@@ -56,7 +56,7 @@
 qed "Fin_UnionI";
 
 (*Every subset of a finite set is finite.*)
-Goal "!!b A. b: Fin(A) ==> ALL z. z<=b --> z: Fin(A)";
+Goal "b: Fin(A) ==> ALL z. z<=b --> z: Fin(A)";
 by (etac Fin_induct 1);
 by (simp_tac (simpset() addsimps [subset_empty_iff]) 1);
 by (asm_simp_tac (simpset() addsimps subset_cons_iff::distrib_simps) 1);
@@ -65,7 +65,7 @@
 by (Asm_simp_tac 1);
 qed "Fin_subset_lemma";
 
-Goal "!!c b A. [| c<=b;  b: Fin(A) |] ==> c: Fin(A)";
+Goal "[| c<=b;  b: Fin(A) |] ==> c: Fin(A)";
 by (REPEAT (ares_tac [Fin_subset_lemma RS spec RS mp] 1));
 qed "Fin_subset";
 
@@ -108,17 +108,17 @@
 by (REPEAT (ares_tac (Fin_mono::Sigma_mono::basic_monos) 1));
 qed "FiniteFun_mono";
 
-Goal "!!A B. A<=B ==> A -||> A  <=  B -||> B";
+Goal "A<=B ==> A -||> A  <=  B -||> B";
 by (REPEAT (ares_tac [FiniteFun_mono] 1));
 qed "FiniteFun_mono1";
 
-Goal "!!h. h: A -||>B ==> h: domain(h) -> B";
+Goal "h: A -||>B ==> h: domain(h) -> B";
 by (etac FiniteFun.induct 1);
 by (simp_tac (simpset() addsimps [empty_fun, domain_0]) 1);
 by (asm_simp_tac (simpset() addsimps [fun_extend3, domain_cons]) 1);
 qed "FiniteFun_is_fun";
 
-Goal "!!h. h: A -||>B ==> domain(h) : Fin(A)";
+Goal "h: A -||>B ==> domain(h) : Fin(A)";
 by (etac FiniteFun.induct 1);
 by (simp_tac (simpset() addsimps [domain_0]) 1);
 by (asm_simp_tac (simpset() addsimps [domain_cons]) 1);
@@ -127,7 +127,7 @@
 bind_thm ("FiniteFun_apply_type", FiniteFun_is_fun RS apply_type);
 
 (*Every subset of a finite function is a finite function.*)
-Goal "!!b A. b: A-||>B ==> ALL z. z<=b --> z: A-||>B";
+Goal "b: A-||>B ==> ALL z. z<=b --> z: A-||>B";
 by (etac FiniteFun.induct 1);
 by (simp_tac (simpset() addsimps subset_empty_iff::FiniteFun.intrs) 1);
 by (asm_simp_tac (simpset() addsimps subset_cons_iff::distrib_simps) 1);
@@ -137,7 +137,7 @@
 by (fast_tac (claset() addSIs FiniteFun.intrs) 1);
 qed "FiniteFun_subset_lemma";
 
-Goal "!!c b A. [| c<=b;  b: A-||>B |] ==> c: A-||>B";
+Goal "[| c<=b;  b: A-||>B |] ==> c: A-||>B";
 by (REPEAT (ares_tac [FiniteFun_subset_lemma RS spec RS mp] 1));
 qed "FiniteFun_subset";
 
--- a/src/ZF/Fixedpt.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Fixedpt.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -38,14 +38,14 @@
 by (rtac subset_refl 1);
 qed "bnd_mono_subset";
 
-Goal "!!A B. [| bnd_mono(D,h);  A <= D;  B <= D |] ==> \
+Goal "[| bnd_mono(D,h);  A <= D;  B <= D |] ==> \
 \                         h(A) Un h(B) <= h(A Un B)";
 by (REPEAT (ares_tac [Un_upper1, Un_upper2, Un_least] 1
      ORELSE etac bnd_monoD2 1));
 qed "bnd_mono_Un";
 
 (*Useful??*)
-Goal "!!A B. [| bnd_mono(D,h);  A <= D;  B <= D |] ==> \
+Goal "[| bnd_mono(D,h);  A <= D;  B <= D |] ==> \
 \                        h(A Int B) <= h(A) Int h(B)";
 by (REPEAT (ares_tac [Int_lower1, Int_lower2, Int_greatest] 1
      ORELSE etac bnd_monoD2 1));
@@ -244,7 +244,7 @@
 (*** Coinduction rules for greatest fixed points ***)
 
 (*weak version*)
-Goal "!!X h. [| a: X;  X <= h(X);  X <= D |] ==> a : gfp(D,h)";
+Goal "[| a: X;  X <= h(X);  X <= D |] ==> a : gfp(D,h)";
 by (REPEAT (ares_tac [gfp_upperbound RS subsetD] 1));
 qed "weak_coinduct";
 
--- a/src/ZF/IMP/Equiv.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/IMP/Equiv.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -44,7 +44,7 @@
 val bexp2 = bexp_iff RS iffD2;
 
 
-Goal "!!c. <c,sigma> -c-> sigma' ==> <sigma,sigma'> : C(c)";
+Goal "<c,sigma> -c-> sigma' ==> <sigma,sigma'> : C(c)";
 by (etac evalc.induct 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [bexp1])));
 (* skip *)
@@ -70,7 +70,7 @@
 AddEs  [C_type,C_type_fst];
 
 
-Goal "!!c. c : com ==> ALL x:C(c). <c,fst(x)> -c-> snd(x)";
+Goal "c : com ==> ALL x:C(c). <c,fst(x)> -c-> snd(x)";
 by (etac com.induct 1);
 (* skip *)
 by (fast_tac (claset() addss (simpset())) 1);
--- a/src/ZF/List.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/List.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -34,7 +34,7 @@
 
 (**  Lemmas to justify using "list" in other recursive type definitions **)
 
-Goalw list.defs "!!A B. A<=B ==> list(A) <= list(B)";
+Goalw list.defs "A<=B ==> list(A) <= list(B)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac list.bnd_mono 1));
 by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
@@ -51,7 +51,7 @@
 (*These two theorems justify datatypes involving list(nat), list(A), ...*)
 bind_thm ("list_subset_univ", ([list_mono, list_univ] MRS subset_trans));
 
-Goal "!!l A B. [| l: list(A);  A <= univ(B) |] ==> l: univ(B)";
+Goal "[| l: list(A);  A <= univ(B) |] ==> l: univ(B)";
 by (REPEAT (ares_tac [list_subset_univ RS subsetD] 1));
 qed "list_into_univ";
 
@@ -94,7 +94,7 @@
 
 Addsimps [hd_Cons, tl_Nil, tl_Cons];
 
-Goal "!!l. l: list(A) ==> tl(l) : list(A)";
+Goal "l: list(A) ==> tl(l) : list(A)";
 by (etac list.elim 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps list.intrs)));
 qed "tl_type";
@@ -105,7 +105,7 @@
 by (rtac rec_0 1);
 qed "drop_0";
 
-Goalw [drop_def] "!!i. i:nat ==> drop(i, Nil) = Nil";
+Goalw [drop_def] "i:nat ==> drop(i, Nil) = Nil";
 by (etac nat_induct 1);
 by (ALLGOALS Asm_simp_tac);
 qed "drop_Nil";
--- a/src/ZF/Nat.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Nat.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -25,10 +25,9 @@
 by (rtac (singletonI RS UnI1) 1);
 qed "nat_0I";
 
-val prems = goal Nat.thy "n : nat ==> succ(n) : nat";
+Goal "n : nat ==> succ(n) : nat";
 by (stac nat_unfold 1);
-by (rtac (RepFunI RS UnI2) 1);
-by (resolve_tac prems 1);
+by (etac (RepFunI RS UnI2) 1);
 qed "nat_succI";
 
 Goal "1 : nat";
@@ -39,8 +38,8 @@
 by (rtac (nat_1I RS nat_succI) 1);
 qed "nat_2I";
 
-Addsimps [nat_0I, nat_1I, nat_2I, nat_succI];
-AddSIs   [nat_0I, nat_1I, nat_2I];
+Addsimps [nat_0I, nat_1I, nat_2I];
+AddSIs   [nat_0I, nat_1I, nat_2I, nat_succI];
 
 Goal "bool <= nat";
 by (REPEAT (ares_tac [subsetI,nat_0I,nat_1I] 1
@@ -93,14 +92,22 @@
 qed "Ord_nat";
 
 Goalw [Limit_def] "Limit(nat)";
-by (safe_tac (claset() addSIs [ltI, nat_succI, Ord_nat]));
+by (safe_tac (claset() addSIs [ltI, Ord_nat]));
 by (etac ltD 1);
 qed "Limit_nat";
 
-Addsimps [Ord_nat, Limit_nat];
+(* succ(i): nat ==> i: nat *)
+val succ_natD = [succI1, asm_rl, Ord_nat] MRS Ord_trans;
+AddSDs   [succ_natD];
+
+Goal "succ(n): nat <-> n: nat";
+by (Blast_tac 1);
+qed "nat_succ_iff";
+
+Addsimps [Ord_nat, Limit_nat, nat_succ_iff];
 AddSIs   [Ord_nat, Limit_nat];
 
-Goal "!!i. Limit(i) ==> nat le i";
+Goal "Limit(i) ==> nat le i";
 by (rtac subset_imp_le 1);
 by (rtac subsetI 1);
 by (etac nat_induct 1);
@@ -109,13 +116,10 @@
                       Ord_nat, Limit_is_Ord] 1));
 qed "nat_le_Limit";
 
-(* succ(i): nat ==> i: nat *)
-val succ_natD = [succI1, asm_rl, Ord_nat] MRS Ord_trans;
-
 (* [| succ(i): k;  k: nat |] ==> i: k *)
 val succ_in_naturalD = [succI1, asm_rl, nat_into_Ord] MRS Ord_trans;
 
-Goal "!!m n. [| m<n;  n: nat |] ==> m: nat";
+Goal "[| m<n;  n: nat |] ==> m: nat";
 by (etac ltE 1);
 by (etac (Ord_nat RSN (3,Ord_trans)) 1);
 by (assume_tac 1);
@@ -221,12 +225,12 @@
 
 (** The union of two natural numbers is a natural number -- their maximum **)
 
-Goal "!!i j. [| i: nat; j: nat |] ==> i Un j: nat";
+Goal "[| i: nat; j: nat |] ==> i Un j: nat";
 by (rtac (Un_least_lt RS ltD) 1);
 by (REPEAT (ares_tac [ltI, Ord_nat] 1));
 qed "Un_nat_type";
 
-Goal "!!i j. [| i: nat; j: nat |] ==> i Int j: nat";
+Goal "[| i: nat; j: nat |] ==> i Int j: nat";
 by (rtac (Int_greatest_lt RS ltD) 1);
 by (REPEAT (ares_tac [ltI, Ord_nat] 1));
 qed "Int_nat_type";
--- a/src/ZF/Order.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Order.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -618,7 +618,7 @@
 (** By Krzysztof Grabczewski.
     Lemmas involving the first element of a well ordered set **)
 
-Goalw [first_def] "!!b. first(b,B,r) ==> b:B";
+Goalw [first_def] "first(b,B,r) ==> b:B";
 by (Blast_tac 1);
 qed "first_is_elem";
 
@@ -633,7 +633,7 @@
 by (Blast.depth_tac (claset()) 7 1);
 qed "well_ord_imp_ex1_first";
 
-Goal "!!r. [| well_ord(A,r); B<=A; B~=0 |] ==> (THE b. first(b,B,r)) : B";
+Goal "[| well_ord(A,r); B<=A; B~=0 |] ==> (THE b. first(b,B,r)) : B";
 by (dtac well_ord_imp_ex1_first 1 THEN REPEAT (assume_tac 1));
 by (rtac first_is_elem 1);
 by (etac theI 1);
--- a/src/ZF/OrderArith.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/OrderArith.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -142,7 +142,7 @@
 by (blast_tac (claset() addSIs [if_type]) 2);
 by (DEPTH_SOLVE_1 (eresolve_tac [case_type, UnI1, UnI2] 1));
 by Safe_tac;
-by (ALLGOALS (asm_simp_tac (simpset() addsplits [split_if])));
+by (ALLGOALS Asm_simp_tac);
 by (blast_tac (claset() addEs [equalityE]) 1);
 qed "sum_disjoint_bij";
 
--- a/src/ZF/OrderType.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/OrderType.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -195,7 +195,7 @@
 (*** Basic equalities for ordertype ***)
 
 (*Ordertype of Memrel*)
-Goal "!!i. j le i ==> ordertype(j,Memrel(i)) = j";
+Goal "j le i ==> ordertype(j,Memrel(i)) = j";
 by (resolve_tac [Ord_iso_implies_eq RS sym] 1);
 by (etac ltE 1);
 by (REPEAT (ares_tac [le_well_ord_Memrel, Ord_ordertype] 1));
@@ -288,7 +288,7 @@
 (**** Alternative definition of ordinal ****)
 
 (*proof by Krzysztof Grabczewski*)
-Goalw [Ord_alt_def] "!!i. Ord(i) ==> Ord_alt(i)";
+Goalw [Ord_alt_def] "Ord(i) ==> Ord_alt(i)";
 by (rtac conjI 1);
 by (etac well_ord_Memrel 1);
 by (rewrite_goals_tac [Ord_def, Transset_def, pred_def, Memrel_def]);
@@ -386,12 +386,12 @@
 
 (** Ordinal addition with zero **)
 
-Goalw [oadd_def] "!!i. Ord(i) ==> i++0 = i";
+Goalw [oadd_def] "Ord(i) ==> i++0 = i";
 by (asm_simp_tac (simpset() addsimps [Memrel_0, ordertype_sum_0_eq, 
                                   ordertype_Memrel, well_ord_Memrel]) 1);
 qed "oadd_0";
 
-Goalw [oadd_def] "!!i. Ord(i) ==> 0++i = i";
+Goalw [oadd_def] "Ord(i) ==> 0++i = i";
 by (asm_simp_tac (simpset() addsimps [Memrel_0, ordertype_0_sum_eq, 
                                   ordertype_Memrel, well_ord_Memrel]) 1);
 qed "oadd_0_left";
@@ -401,7 +401,7 @@
 (*** Further properties of ordinal addition.  Statements by Grabczewski,
     proofs by lcp. ***)
 
-Goalw [oadd_def] "!!i j k. [| k<i;  Ord(j) |] ==> k < i++j";
+Goalw [oadd_def] "[| k<i;  Ord(j) |] ==> k < i++j";
 by (rtac ltE 1 THEN assume_tac 1);
 by (rtac ltI 1);
 by (REPEAT (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel] 2));
@@ -414,7 +414,7 @@
 qed "lt_oadd1";
 
 (*Thus also we obtain the rule  i++j = k ==> i le k *)
-Goal "!!i j. [| Ord(i);  Ord(j) |] ==> i le i++j";
+Goal "[| Ord(i);  Ord(j) |] ==> i le i++j";
 by (rtac all_lt_imp_le 1);
 by (REPEAT (ares_tac [Ord_oadd, lt_oadd1] 1));
 qed "oadd_le_self";
@@ -438,7 +438,7 @@
 by (REPEAT_FIRST (ares_tac [well_ord_radd, well_ord_Memrel]));
 qed "ordertype_sum_Memrel";
 
-Goalw [oadd_def] "!!i j k. [| k<j;  Ord(i) |] ==> i++k < i++j";
+Goalw [oadd_def] "[| k<j;  Ord(i) |] ==> i++k < i++j";
 by (rtac ltE 1 THEN assume_tac 1);
 by (resolve_tac [ordertype_pred_unfold RS equalityD2 RS subsetD RS ltI] 1);
 by (REPEAT_FIRST (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel]));
@@ -513,7 +513,7 @@
 by (blast_tac (claset() addSEs [ltE]) 1);
 qed "oadd_unfold";
 
-Goal "!!i. Ord(i) ==> i++1 = succ(i)";
+Goal "Ord(i) ==> i++1 = succ(i)";
 by (asm_simp_tac (simpset() addsimps [oadd_unfold, Ord_1, oadd_0]) 1);
 by (Blast_tac 1);
 qed "oadd_1";
@@ -546,7 +546,7 @@
 
 (** Order/monotonicity properties of ordinal addition **)
 
-Goal "!!i j. [| Ord(i);  Ord(j) |] ==> i le j++i";
+Goal "[| Ord(i);  Ord(j) |] ==> i le j++i";
 by (eres_inst_tac [("i","i")] trans_induct3 1);
 by (asm_simp_tac (simpset() addsimps [Ord_0_le]) 1);
 by (asm_simp_tac (simpset() addsimps [oadd_succ, succ_leI]) 1);
@@ -558,7 +558,7 @@
                                      le_refl, Limit_is_Ord]) 1);
 qed "oadd_le_self2";
 
-Goal "!!i j k. [| k le j;  Ord(i) |] ==> k++i le j++i";
+Goal "[| k le j;  Ord(i) |] ==> k++i le j++i";
 by (forward_tac [lt_Ord] 1);
 by (forward_tac [le_Ord2] 1);
 by (etac trans_induct3 1);
@@ -569,13 +569,13 @@
 by (Blast_tac 1);
 qed "oadd_le_mono1";
 
-Goal "!!i j. [| i' le i;  j'<j |] ==> i'++j' < i++j";
+Goal "[| i' le i;  j'<j |] ==> i'++j' < i++j";
 by (rtac lt_trans1 1);
 by (REPEAT (eresolve_tac [asm_rl, oadd_le_mono1, oadd_lt_mono2, ltE,
                           Ord_succD] 1));
 qed "oadd_lt_mono";
 
-Goal "!!i j. [| i' le i;  j' le j |] ==> i'++j' le i++j";
+Goal "[| i' le i;  j' le j |] ==> i'++j' le i++j";
 by (asm_simp_tac (simpset() addsimps [oadd_succ RS sym, le_Ord2, oadd_lt_mono]) 1);
 qed "oadd_le_mono";
 
@@ -596,7 +596,7 @@
 by (blast_tac (claset() addSIs [if_type]) 1);
 by (fast_tac (claset() addSIs [case_type]) 1);
 by (etac sumE 2);
-by (ALLGOALS (asm_simp_tac (simpset() addsplits [split_if])));
+by (ALLGOALS Asm_simp_tac);
 qed "bij_sum_Diff";
 
 Goal
@@ -607,8 +607,7 @@
 by (resolve_tac [bij_sum_Diff RS ord_isoI RS ord_iso_sym RS ordertype_eq] 1);
 by (etac well_ord_Memrel 3);
 by (assume_tac 1);
-by (asm_simp_tac 
-     (simpset() addsplits [split_if] addsimps [Memrel_iff]) 1);
+by (asm_simp_tac (simpset() addsimps [Memrel_iff]) 1);
 by (forw_inst_tac [("j", "y")] Ord_in_Ord 1 THEN assume_tac 1);
 by (forw_inst_tac [("j", "x")] Ord_in_Ord 1 THEN assume_tac 1);
 by (asm_simp_tac (simpset() addsimps [Ord_mem_iff_lt, lt_Ord, not_lt_iff_le]) 1);
@@ -626,7 +625,7 @@
                       Diff_subset] 1));
 qed "oadd_ordertype_Diff";
 
-Goal "!!i j. i le j ==> i ++ (j--i) = j";
+Goal "i le j ==> i ++ (j--i) = j";
 by (asm_simp_tac (simpset() addsimps [oadd_ordertype_Diff, ordertype_sum_Diff, 
                                   ordertype_Memrel, lt_Ord2 RS Ord_succD]) 1);
 qed "oadd_odiff_inverse";
@@ -760,7 +759,7 @@
 
 (** Ordinal multiplication by 1 **)
 
-Goalw [omult_def] "!!i. Ord(i) ==> i**1 = i";
+Goalw [omult_def] "Ord(i) ==> i**1 = i";
 by (resolve_tac [ord_isoI RS ordertype_eq RS trans] 1);
 by (res_inst_tac [("c", "snd"), ("d", "%z.<0,z>")] lam_bijective 1);
 by (REPEAT_FIRST (eresolve_tac [snd_type, SigmaE, succE, emptyE, 
@@ -768,7 +767,7 @@
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [rmult_iff, Memrel_iff])));
 qed "omult_1";
 
-Goalw [omult_def] "!!i. Ord(i) ==> 1**i = i";
+Goalw [omult_def] "Ord(i) ==> 1**i = i";
 by (resolve_tac [ord_isoI RS ordertype_eq RS trans] 1);
 by (res_inst_tac [("c", "fst"), ("d", "%z.<z,0>")] lam_bijective 1);
 by (REPEAT_FIRST (eresolve_tac [fst_type, SigmaE, succE, emptyE, 
@@ -794,7 +793,7 @@
                       Ord_ordertype] 1));
 qed "oadd_omult_distrib";
 
-Goal "!!i. [| Ord(i);  Ord(j) |] ==> i**succ(j) = (i**j)++i";
+Goal "[| Ord(i);  Ord(j) |] ==> i**succ(j) = (i**j)++i";
                 (*ZF_ss prevents looping*)
 by (asm_simp_tac (ZF_ss addsimps [oadd_1 RS sym]) 1);
 by (asm_simp_tac 
@@ -837,19 +836,19 @@
 (*** Ordering/monotonicity properties of ordinal multiplication ***)
 
 (*As a special case we have "[| 0<i;  0<j |] ==> 0 < i**j" *)
-Goal "!!i j. [| k<i;  0<j |] ==> k < i**j";
+Goal "[| k<i;  0<j |] ==> k < i**j";
 by (safe_tac (claset() addSEs [ltE] addSIs [ltI, Ord_omult]));
 by (asm_simp_tac (simpset() addsimps [omult_unfold]) 1);
 by (REPEAT_FIRST (ares_tac [bexI]));
 by (Asm_simp_tac 1);
 qed "lt_omult1";
 
-Goal "!!i j. [| Ord(i);  0<j |] ==> i le i**j";
+Goal "[| Ord(i);  0<j |] ==> i le i**j";
 by (rtac all_lt_imp_le 1);
 by (REPEAT (ares_tac [Ord_omult, lt_omult1, lt_Ord2] 1));
 qed "omult_le_self";
 
-Goal "!!i j k. [| k le j;  Ord(i) |] ==> k**i le j**i";
+Goal "[| k le j;  Ord(i) |] ==> k**i le j**i";
 by (forward_tac [lt_Ord] 1);
 by (forward_tac [le_Ord2] 1);
 by (etac trans_induct3 1);
@@ -860,7 +859,7 @@
 by (Blast_tac 1);
 qed "omult_le_mono1";
 
-Goal "!!i j k. [| k<j;  0<i |] ==> i**k < i**j";
+Goal "[| k<j;  0<i |] ==> i**k < i**j";
 by (rtac ltI 1);
 by (asm_simp_tac (simpset() addsimps [omult_unfold, lt_Ord2]) 1);
 by (safe_tac (claset() addSEs [ltE] addSIs [Ord_omult]));
@@ -868,14 +867,14 @@
 by (asm_simp_tac (simpset() addsimps [Ord_omult]) 1);
 qed "omult_lt_mono2";
 
-Goal "!!i j k. [| k le j;  Ord(i) |] ==> i**k le i**j";
+Goal "[| k le j;  Ord(i) |] ==> i**k le i**j";
 by (rtac subset_imp_le 1);
 by (safe_tac (claset() addSEs [ltE, make_elim Ord_succD] addSIs [Ord_omult]));
 by (asm_full_simp_tac (simpset() addsimps [omult_unfold]) 1);
 by (deepen_tac (claset() addEs [Ord_trans]) 0 1);
 qed "omult_le_mono2";
 
-Goal "!!i j. [| i' le i;  j' le j |] ==> i'**j' le i**j";
+Goal "[| i' le i;  j' le j |] ==> i'**j' le i**j";
 by (rtac le_trans 1);
 by (REPEAT (eresolve_tac [asm_rl, omult_le_mono1, omult_le_mono2, ltE,
                           Ord_succD] 1));
@@ -888,7 +887,7 @@
                           Ord_succD] 1));
 qed "omult_lt_mono";
 
-Goal "!!i j. [| Ord(i);  0<j |] ==> i le j**i";
+Goal "[| Ord(i);  0<j |] ==> i le j**i";
 by (forward_tac [lt_Ord2] 1);
 by (eres_inst_tac [("i","i")] trans_induct3 1);
 by (asm_simp_tac (simpset() addsimps [omult_0, Ord_0 RS le_refl]) 1);
@@ -908,7 +907,7 @@
 
 (** Further properties of ordinal multiplication **)
 
-Goal "!!i j. [| i**j = i**k;  0<i;  Ord(j);  Ord(k) |] ==> j=k";
+Goal "[| i**j = i**k;  0<i;  Ord(j);  Ord(k) |] ==> j=k";
 by (rtac Ord_linear_lt 1);
 by (REPEAT_SOME assume_tac);
 by (ALLGOALS
--- a/src/ZF/Ordinal.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Ordinal.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -61,15 +61,15 @@
 by (Blast_tac 1);
 qed "Transset_Int";
 
-Goalw [Transset_def] "!!i. Transset(i) ==> Transset(succ(i))";
+Goalw [Transset_def] "Transset(i) ==> Transset(succ(i))";
 by (Blast_tac 1);
 qed "Transset_succ";
 
-Goalw [Transset_def] "!!i. Transset(i) ==> Transset(Pow(i))";
+Goalw [Transset_def] "Transset(i) ==> Transset(Pow(i))";
 by (Blast_tac 1);
 qed "Transset_Pow";
 
-Goalw [Transset_def] "!!A. Transset(A) ==> Transset(Union(A))";
+Goalw [Transset_def] "Transset(A) ==> Transset(Union(A))";
 by (Blast_tac 1);
 qed "Transset_Union";
 
@@ -112,7 +112,7 @@
 
 AddSDs [Ord_succD];
 
-Goal "!!i j. [| Ord(i);  Transset(j);  j<=i |] ==> Ord(j)";
+Goal "[| Ord(i);  Transset(j);  j<=i |] ==> Ord(j)";
 by (REPEAT (ares_tac [OrdI] 1
      ORELSE eresolve_tac [Ord_contains_Transset, subsetD] 1));
 qed "Ord_subset_Ord";
@@ -121,11 +121,11 @@
 by (Blast_tac 1);
 qed "OrdmemD";
 
-Goal "!!i j k. [| i:j;  j:k;  Ord(k) |] ==> i:k";
+Goal "[| i:j;  j:k;  Ord(k) |] ==> i:k";
 by (REPEAT (ares_tac [OrdmemD RS subsetD] 1));
 qed "Ord_trans";
 
-Goal "!!i j. [| i:j;  Ord(j) |] ==> succ(i) <= j";
+Goal "[| i:j;  Ord(j) |] ==> succ(i) <= j";
 by (REPEAT (ares_tac [OrdmemD RSN (2,succ_subsetI)] 1));
 qed "Ord_succ_subsetI";
 
@@ -136,7 +136,7 @@
 by (REPEAT (ares_tac [OrdI,Transset_0] 1 ORELSE etac emptyE 1));
 qed "Ord_0";
 
-Goal "!!i. Ord(i) ==> Ord(succ(i))";
+Goal "Ord(i) ==> Ord(succ(i))";
 by (REPEAT (ares_tac [OrdI,Transset_succ] 1
      ORELSE eresolve_tac [succE,ssubst,Ord_is_Transset,
                           Ord_contains_Transset] 1));
@@ -151,11 +151,11 @@
 Addsimps [Ord_0, Ord_succ_iff];
 AddSIs   [Ord_0, Ord_succ];
 
-Goalw [Ord_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i Un j)";
+Goalw [Ord_def] "[| Ord(i); Ord(j) |] ==> Ord(i Un j)";
 by (blast_tac (claset() addSIs [Transset_Un]) 1);
 qed "Ord_Un";
 
-Goalw [Ord_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i Int j)";
+Goalw [Ord_def] "[| Ord(i); Ord(j) |] ==> Ord(i Int j)";
 by (blast_tac (claset() addSIs [Transset_Int]) 1);
 qed "Ord_Int";
 
@@ -190,7 +190,7 @@
 
 (*** < is 'less than' for ordinals ***)
 
-Goalw [lt_def] "!!i j. [| i:j;  Ord(j) |] ==> i<j";
+Goalw [lt_def] "[| i:j;  Ord(j) |] ==> i<j";
 by (REPEAT (ares_tac [conjI] 1));
 qed "ltI";
 
@@ -200,7 +200,7 @@
 by (REPEAT (ares_tac (prems@[Ord_in_Ord]) 1));
 qed "ltE";
 
-Goal "!!i j. i<j ==> i:j";
+Goal "i<j ==> i:j";
 by (etac ltE 1);
 by (assume_tac 1);
 qed "ltD";
@@ -211,11 +211,11 @@
 
 Addsimps [not_lt0];
 
-Goal "!!i j. j<i ==> Ord(j)";
+Goal "j<i ==> Ord(j)";
 by (etac ltE 1 THEN assume_tac 1);
 qed "lt_Ord";
 
-Goal "!!i j. j<i ==> Ord(i)";
+Goal "j<i ==> Ord(i)";
 by (etac ltE 1 THEN assume_tac 1);
 qed "lt_Ord2";
 
@@ -225,11 +225,11 @@
 (* i<0 ==> R *)
 bind_thm ("lt0E", not_lt0 RS notE);
 
-Goal "!!i j k. [| i<j;  j<k |] ==> i<k";
+Goal "[| i<j;  j<k |] ==> i<k";
 by (blast_tac (claset() addSIs [ltI] addSEs [ltE] addIs [Ord_trans]) 1);
 qed "lt_trans";
 
-Goalw [lt_def] "!!i j. [| i<j;  j<i |] ==> P";
+Goalw [lt_def] "[| i<j;  j<i |] ==> P";
 by (REPEAT (eresolve_tac [asm_rl, conjE, mem_asym] 1));
 qed "lt_asym";
 
@@ -248,11 +248,11 @@
 qed "le_iff";
 
 (*Equivalently, i<j ==> i < succ(j)*)
-Goal "!!i j. i<j ==> i le j";
+Goal "i<j ==> i le j";
 by (asm_simp_tac (simpset() addsimps [le_iff]) 1);
 qed "leI";
 
-Goal "!!i. [| i=j;  Ord(j) |] ==> i le j";
+Goal "[| i=j;  Ord(j) |] ==> i le j";
 by (asm_simp_tac (simpset() addsimps [le_iff]) 1);
 qed "le_eqI";
 
@@ -269,7 +269,7 @@
 by (DEPTH_SOLVE (ares_tac prems 1 ORELSE etac conjE 1));
 qed "leE";
 
-Goal "!!i j. [| i le j;  j le i |] ==> i=j";
+Goal "[| i le j;  j le i |] ==> i=j";
 by (asm_full_simp_tac (simpset() addsimps [le_iff]) 1);
 by (blast_tac (claset() addEs [lt_asym]) 1);
 qed "le_anti_sym";
@@ -294,7 +294,7 @@
 by (Blast_tac 1);
 qed "Memrel_iff";
 
-Goal "!!A. [| a: b;  a: A;  b: A |] ==> <a,b> : Memrel(A)";
+Goal "[| a: b;  a: A;  b: A |] ==> <a,b> : Memrel(A)";
 by (REPEAT (ares_tac [conjI, Memrel_iff RS iffD2] 1));
 qed "MemrelI";
 
@@ -315,7 +315,7 @@
 by (Blast_tac 1);
 qed "Memrel_type";
 
-Goalw [Memrel_def] "!!A B. A<=B ==> Memrel(A) <= Memrel(B)";
+Goalw [Memrel_def] "A<=B ==> Memrel(A) <= Memrel(B)";
 by (Blast_tac 1);
 qed "Memrel_mono";
 
@@ -425,11 +425,11 @@
 by (DEPTH_SOLVE (ares_tac ([leI,le_eqI] @ prems) 1));
 qed "Ord_linear_le";
 
-Goal "!!i j. j le i ==> ~ i<j";
+Goal "j le i ==> ~ i<j";
 by (blast_tac le_cs 1);
 qed "le_imp_not_lt";
 
-Goal "!!i j. [| ~ i<j;  Ord(i);  Ord(j) |] ==> j le i";
+Goal "[| ~ i<j;  Ord(i);  Ord(j) |] ==> j le i";
 by (res_inst_tac [("i","i"),("j","j")] Ord_linear2 1);
 by (REPEAT (SOMEGOAL assume_tac));
 by (blast_tac le_cs 1);
@@ -437,25 +437,25 @@
 
 (** Some rewrite rules for <, le **)
 
-Goalw [lt_def] "!!i j. Ord(j) ==> i:j <-> i<j";
+Goalw [lt_def] "Ord(j) ==> i:j <-> i<j";
 by (Blast_tac 1);
 qed "Ord_mem_iff_lt";
 
-Goal "!!i j. [| Ord(i);  Ord(j) |] ==> ~ i<j <-> j le i";
+Goal "[| Ord(i);  Ord(j) |] ==> ~ i<j <-> j le i";
 by (REPEAT (ares_tac [iffI, le_imp_not_lt, not_lt_imp_le] 1));
 qed "not_lt_iff_le";
 
-Goal "!!i j. [| Ord(i);  Ord(j) |] ==> ~ i le j <-> j<i";
+Goal "[| Ord(i);  Ord(j) |] ==> ~ i le j <-> j<i";
 by (asm_simp_tac (simpset() addsimps [not_lt_iff_le RS iff_sym]) 1);
 qed "not_le_iff_lt";
 
 (*This is identical to 0<succ(i) *)
-Goal "!!i. Ord(i) ==> 0 le i";
+Goal "Ord(i) ==> 0 le i";
 by (etac (not_lt_iff_le RS iffD1) 1);
 by (REPEAT (resolve_tac [Ord_0, not_lt0] 1));
 qed "Ord_0_le";
 
-Goal "!!i. [| Ord(i);  i~=0 |] ==> 0<i";
+Goal "[| Ord(i);  i~=0 |] ==> 0<i";
 by (etac (not_le_iff_lt RS iffD1) 1);
 by (rtac Ord_0 1);
 by (Blast_tac 1);
@@ -465,14 +465,14 @@
 
 (** For ordinals, j<=i (subset) implies j le i (less-than or equals) **)
 
-Goal "!!i j. [| j<=i;  Ord(i);  Ord(j) |] ==> j le i";
+Goal "[| j<=i;  Ord(i);  Ord(j) |] ==> j le i";
 by (rtac (not_lt_iff_le RS iffD1) 1);
 by (assume_tac 1);
 by (assume_tac 1);
 by (blast_tac (claset() addEs [ltE, mem_irrefl]) 1);
 qed "subset_imp_le";
 
-Goal "!!i j. i le j ==> i<=j";
+Goal "i le j ==> i<=j";
 by (etac leE 1);
 by (Blast_tac 2);
 by (blast_tac (subset_cs addIs [OrdmemD] addEs [ltE]) 1);
@@ -497,26 +497,26 @@
 
 (** Transitive laws **)
 
-Goal "!!i j. [| i le j;  j<k |] ==> i<k";
+Goal "[| i le j;  j<k |] ==> i<k";
 by (blast_tac (claset() addSEs [leE] addIs [lt_trans]) 1);
 qed "lt_trans1";
 
-Goal "!!i j. [| i<j;  j le k |] ==> i<k";
+Goal "[| i<j;  j le k |] ==> i<k";
 by (blast_tac (claset() addSEs [leE] addIs [lt_trans]) 1);
 qed "lt_trans2";
 
-Goal "!!i j. [| i le j;  j le k |] ==> i le k";
+Goal "[| i le j;  j le k |] ==> i le k";
 by (REPEAT (ares_tac [lt_trans1] 1));
 qed "le_trans";
 
-Goal "!!i j. i<j ==> succ(i) le j";
+Goal "i<j ==> succ(i) le j";
 by (rtac (not_lt_iff_le RS iffD1) 1);
 by (blast_tac le_cs 3);
 by (ALLGOALS (blast_tac (claset() addEs [ltE])));
 qed "succ_leI";
 
 (*Identical to  succ(i) < succ(j) ==> i<j  *)
-Goal "!!i j. succ(i) le j ==> i<j";
+Goal "succ(i) le j ==> i<j";
 by (rtac (not_le_iff_lt RS iffD1) 1);
 by (blast_tac le_cs 3);
 by (ALLGOALS (blast_tac (claset() addEs [ltE, make_elim Ord_succD])));
@@ -528,24 +528,24 @@
 
 Addsimps [succ_le_iff];
 
-Goal "!!i j. succ(i) le succ(j) ==> i le j";
+Goal "succ(i) le succ(j) ==> i le j";
 by (blast_tac (claset() addSDs [succ_leE]) 1);
 qed "succ_le_imp_le";
 
 (** Union and Intersection **)
 
-Goal "!!i j. [| Ord(i); Ord(j) |] ==> i le i Un j";
+Goal "[| Ord(i); Ord(j) |] ==> i le i Un j";
 by (rtac (Un_upper1 RS subset_imp_le) 1);
 by (REPEAT (ares_tac [Ord_Un] 1));
 qed "Un_upper1_le";
 
-Goal "!!i j. [| Ord(i); Ord(j) |] ==> j le i Un j";
+Goal "[| Ord(i); Ord(j) |] ==> j le i Un j";
 by (rtac (Un_upper2 RS subset_imp_le) 1);
 by (REPEAT (ares_tac [Ord_Un] 1));
 qed "Un_upper2_le";
 
 (*Replacing k by succ(k') yields the similar rule for le!*)
-Goal "!!i j k. [| i<k;  j<k |] ==> i Un j < k";
+Goal "[| i<k;  j<k |] ==> i Un j < k";
 by (res_inst_tac [("i","i"),("j","j")] Ord_linear_le 1);
 by (stac Un_commute 4);
 by (asm_full_simp_tac (simpset() addsimps [le_subset_iff, subset_Un_iff]) 4);
@@ -553,7 +553,7 @@
 by (REPEAT (eresolve_tac [asm_rl, ltE] 1));
 qed "Un_least_lt";
 
-Goal "!!i j. [| Ord(i); Ord(j) |] ==> i Un j < k  <->  i<k & j<k";
+Goal "[| Ord(i); Ord(j) |] ==> i Un j < k  <->  i<k & j<k";
 by (safe_tac (claset() addSIs [Un_least_lt]));
 by (rtac (Un_upper2_le RS lt_trans1) 2);
 by (rtac (Un_upper1_le RS lt_trans1) 1);
@@ -568,7 +568,7 @@
 qed "Un_least_mem_iff";
 
 (*Replacing k by succ(k') yields the similar rule for le!*)
-Goal "!!i j k. [| i<k;  j<k |] ==> i Int j < k";
+Goal "[| i<k;  j<k |] ==> i Int j < k";
 by (res_inst_tac [("i","i"),("j","j")] Ord_linear_le 1);
 by (stac Int_commute 4);
 by (asm_full_simp_tac (simpset() addsimps [le_subset_iff, subset_Int_iff]) 4);
@@ -625,31 +625,31 @@
      ORELSE dtac Ord_succD 1));
 qed "le_implies_UN_le_UN";
 
-Goal "!!i. Ord(i) ==> (UN y:i. succ(y)) = i";
+Goal "Ord(i) ==> (UN y:i. succ(y)) = i";
 by (blast_tac (claset() addIs [Ord_trans]) 1);
 qed "Ord_equality";
 
 (*Holds for all transitive sets, not just ordinals*)
-Goal "!!i. Ord(i) ==> Union(i) <= i";
+Goal "Ord(i) ==> Union(i) <= i";
 by (blast_tac (claset() addIs [Ord_trans]) 1);
 qed "Ord_Union_subset";
 
 
 (*** Limit ordinals -- general properties ***)
 
-Goalw [Limit_def] "!!i. Limit(i) ==> Union(i) = i";
+Goalw [Limit_def] "Limit(i) ==> Union(i) = i";
 by (fast_tac (claset() addSIs [ltI] addSEs [ltE] addEs [Ord_trans]) 1);
 qed "Limit_Union_eq";
 
-Goalw [Limit_def] "!!i. Limit(i) ==> Ord(i)";
+Goalw [Limit_def] "Limit(i) ==> Ord(i)";
 by (etac conjunct1 1);
 qed "Limit_is_Ord";
 
-Goalw [Limit_def] "!!i. Limit(i) ==> 0 < i";
+Goalw [Limit_def] "Limit(i) ==> 0 < i";
 by (etac (conjunct2 RS conjunct1) 1);
 qed "Limit_has_0";
 
-Goalw [Limit_def] "!!i. [| Limit(i);  j<i |] ==> succ(j) < i";
+Goalw [Limit_def] "[| Limit(i);  j<i |] ==> succ(j) < i";
 by (Blast_tac 1);
 qed "Limit_has_succ";
 
@@ -661,20 +661,20 @@
 by (REPEAT (eresolve_tac [asm_rl, ltE, Ord_succ] 1));
 qed "non_succ_LimitI";
 
-Goal "!!i. Limit(succ(i)) ==> P";
+Goal "Limit(succ(i)) ==> P";
 by (rtac lt_irrefl 1);
 by (rtac Limit_has_succ 1);
 by (assume_tac 1);
 by (etac (Limit_is_Ord RS Ord_succD RS le_refl) 1);
 qed "succ_LimitE";
 
-Goal "!!i. [| Limit(i);  i le succ(j) |] ==> i le j";
+Goal "[| Limit(i);  i le succ(j) |] ==> i le j";
 by (safe_tac (claset() addSEs [succ_LimitE, leE]));
 qed "Limit_le_succD";
 
 (** Traditional 3-way case analysis on ordinals **)
 
-Goal "!!i. Ord(i) ==> i=0 | (EX j. Ord(j) & i=succ(j)) | Limit(i)";
+Goal "Ord(i) ==> i=0 | (EX j. Ord(j) & i=succ(j)) | Limit(i)";
 by (blast_tac (claset() addSIs [non_succ_LimitI, Ord_0_lt]
                       addSDs [Ord_succD]) 1);
 qed "Ord_cases_disj";
--- a/src/ZF/Perm.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Perm.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -13,15 +13,15 @@
 
 (** Surjective function space **)
 
-Goalw [surj_def] "!!f A B. f: surj(A,B) ==> f: A->B";
+Goalw [surj_def] "f: surj(A,B) ==> f: A->B";
 by (etac CollectD1 1);
 qed "surj_is_fun";
 
-Goalw [surj_def] "!!f A B. f : Pi(A,B) ==> f: surj(A,range(f))";
+Goalw [surj_def] "f : Pi(A,B) ==> f: surj(A,range(f))";
 by (blast_tac (claset() addIs [apply_equality, range_of_fun, domain_type]) 1);
 qed "fun_is_surj";
 
-Goalw [surj_def] "!!f A B. f: surj(A,B) ==> range(f)=B";
+Goalw [surj_def] "f: surj(A,B) ==> range(f)=B";
 by (best_tac (claset() addIs [apply_Pair] addEs [range_type]) 1);
 qed "surj_range";
 
@@ -52,7 +52,7 @@
 
 (** Injective function space **)
 
-Goalw [inj_def] "!!f A B. f: inj(A,B) ==> f: A->B";
+Goalw [inj_def] "f: inj(A,B) ==> f: A->B";
 by (etac CollectD1 1);
 qed "inj_is_fun";
 
@@ -63,13 +63,13 @@
 by (Blast_tac 1);
 qed "inj_equality";
 
-Goalw [inj_def] "!!A B f. [| f:inj(A,B);  a:A;  b:A;  f`a=f`b |] ==> a=b";
+Goalw [inj_def] "[| f:inj(A,B);  a:A;  b:A;  f`a=f`b |] ==> a=b";
 by (Blast_tac 1);
 val inj_apply_equality = result();
 
 (** A function with a left inverse is an injection **)
 
-Goal "!!f. [| f: A->B;  ALL x:A. d(f`x)=x |] ==> f: inj(A,B)";
+Goal "[| f: A->B;  ALL x:A. d(f`x)=x |] ==> f: inj(A,B)";
 by (asm_simp_tac (simpset() addsimps [inj_def]) 1);
 by (deepen_tac (claset() addEs [subst_context RS box_equals]) 0 1);
 bind_thm ("f_imp_injective", ballI RSN (2,result()));
@@ -84,11 +84,11 @@
 
 (** Bijections **)
 
-Goalw [bij_def] "!!f A B. f: bij(A,B) ==> f: inj(A,B)";
+Goalw [bij_def] "f: bij(A,B) ==> f: inj(A,B)";
 by (etac IntD1 1);
 qed "bij_is_inj";
 
-Goalw [bij_def] "!!f A B. f: bij(A,B) ==> f: surj(A,B)";
+Goalw [bij_def] "f: bij(A,B) ==> f: surj(A,B)";
 by (etac IntD2 1);
 qed "bij_is_surj";
 
@@ -123,7 +123,7 @@
 by (assume_tac 1);
 qed "id_type";
 
-Goalw [id_def] "!!A x. x:A ==> id(A)`x = x";
+Goalw [id_def] "x:A ==> id(A)`x = x";
 by (Asm_simp_tac 1);
 qed "id_conv";
 
@@ -157,7 +157,7 @@
 
 (*** Converse of a function ***)
 
-Goalw [inj_def] "!!f. f: inj(A,B) ==> converse(f) : range(f)->A";
+Goalw [inj_def] "f: inj(A,B) ==> converse(f) : range(f)->A";
 by (asm_simp_tac (simpset() addsimps [Pi_iff, function_def]) 1);
 by (etac CollectE 1);
 by (asm_simp_tac (simpset() addsimps [apply_iff]) 1);
@@ -188,7 +188,7 @@
 
 (*Should the premises be f:surj(A,B), b:B for symmetry with left_inverse?
   No: they would not imply that converse(f) was a function! *)
-Goal "!!f. [| f: inj(A,B);  b: range(f) |] ==> f`(converse(f)`b) = b";
+Goal "[| f: inj(A,B);  b: range(f) |] ==> f`(converse(f)`b) = b";
 by (rtac right_inverse_lemma 1);
 by (REPEAT (ares_tac [inj_converse_fun,inj_is_fun] 1));
 qed "right_inverse";
@@ -196,26 +196,26 @@
 (*Cannot add [left_inverse, right_inverse] to default simpset: there are too
   many ways of expressing sufficient conditions.*)
 
-Goal "!!f. [| f: bij(A,B);  b: B |] ==> f`(converse(f)`b) = b";
+Goal "[| f: bij(A,B);  b: B |] ==> f`(converse(f)`b) = b";
 by (fast_tac (claset() addss
 	      (simpset() addsimps [bij_def, right_inverse, surj_range])) 1);
 qed "right_inverse_bij";
 
 (** Converses of injections, surjections, bijections **)
 
-Goal "!!f A B. f: inj(A,B) ==> converse(f): inj(range(f), A)";
+Goal "f: inj(A,B) ==> converse(f): inj(range(f), A)";
 by (rtac f_imp_injective 1);
 by (etac inj_converse_fun 1);
 by (rtac right_inverse 1);
 by (REPEAT (assume_tac 1));
 qed "inj_converse_inj";
 
-Goal "!!f A B. f: inj(A,B) ==> converse(f): surj(range(f), A)";
+Goal "f: inj(A,B) ==> converse(f): surj(range(f), A)";
 by (blast_tac (claset() addIs [f_imp_surjective, inj_converse_fun, left_inverse,
 			      inj_is_fun, range_of_fun RS apply_type]) 1);
 qed "inj_converse_surj";
 
-Goalw [bij_def] "!!f A B. f: bij(A,B) ==> converse(f): bij(B,A)";
+Goalw [bij_def] "f: bij(A,B) ==> converse(f): bij(B,A)";
 by (fast_tac (claset() addEs [surj_range RS subst, inj_converse_inj,
 			      inj_converse_surj]) 1);
 qed "bij_converse_bij";
@@ -226,7 +226,7 @@
 
 (*The inductive definition package could derive these theorems for (r O s)*)
 
-Goalw [comp_def] "!!r s. [| <a,b>:s; <b,c>:r |] ==> <a,c> : r O s";
+Goalw [comp_def] "[| <a,b>:s; <b,c>:r |] ==> <a,c> : r O s";
 by (Blast_tac 1);
 qed "compI";
 
@@ -254,7 +254,7 @@
 by (Blast_tac 1);
 qed "range_comp";
 
-Goal "!!r s. domain(r) <= range(s) ==> range(r O s) = range(r)";
+Goal "domain(r) <= range(s) ==> range(r O s) = range(r)";
 by (rtac (range_comp RS equalityI) 1);
 by (Blast_tac 1);
 qed "range_comp_eq";
@@ -263,7 +263,7 @@
 by (Blast_tac 1);
 qed "domain_comp";
 
-Goal "!!r s. range(s) <= domain(r) ==> domain(r O s) = domain(s)";
+Goal "range(s) <= domain(r) ==> domain(r O s) = domain(s)";
 by (rtac (domain_comp RS equalityI) 1);
 by (Blast_tac 1);
 qed "domain_comp_eq";
@@ -275,12 +275,12 @@
 
 (** Other results **)
 
-Goal "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)";
+Goal "[| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)";
 by (Blast_tac 1);
 qed "comp_mono";
 
 (*composition preserves relations*)
-Goal "!!r s. [| s<=A*B;  r<=B*C |] ==> (r O s) <= A*C";
+Goal "[| s<=A*B;  r<=B*C |] ==> (r O s) <= A*C";
 by (Blast_tac 1);
 qed "comp_rel";
 
@@ -292,14 +292,14 @@
 (*left identity of composition; provable inclusions are
         id(A) O r <= r       
   and   [| r<=A*B; B<=C |] ==> r <= id(C) O r *)
-Goal "!!r A B. r<=A*B ==> id(B) O r = r";
+Goal "r<=A*B ==> id(B) O r = r";
 by (Blast_tac 1);
 qed "left_comp_id";
 
 (*right identity of composition; provable inclusions are
         r O id(A) <= r
   and   [| r<=A*B; A<=C |] ==> r <= r O id(C) *)
-Goal "!!r A B. r<=A*B ==> r O id(A) = r";
+Goal "r<=A*B ==> r O id(A) = r";
 by (Blast_tac 1);
 qed "right_comp_id";
 
@@ -311,7 +311,7 @@
 by (Blast_tac 1);
 qed "comp_function";
 
-Goal "!!f g. [| g: A->B;  f: B->C |] ==> (f O g) : A->C";
+Goal "[| g: A->B;  f: B->C |] ==> (f O g) : A->C";
 by (asm_full_simp_tac
     (simpset() addsimps [Pi_def, comp_function, Pow_iff, comp_rel]
            setloop etac conjE) 1);
@@ -319,7 +319,7 @@
 by (Blast_tac 1);
 qed "comp_fun";
 
-Goal "!!f g. [| g: A->B;  f: B->C;  a:A |] ==> (f O g)`a = f`(g`a)";
+Goal "[| g: A->B;  f: B->C;  a:A |] ==> (f O g)`a = f`(g`a)";
 by (REPEAT (ares_tac [comp_fun,apply_equality,compI,
                       apply_Pair,apply_type] 1));
 qed "comp_fun_apply";
@@ -338,7 +338,7 @@
              setSolver type_auto_tac [lam_type, lam_funtype, prem]) 1);
 qed "comp_lam";
 
-Goal "!!f g. [| g: inj(A,B);  f: inj(B,C) |] ==> (f O g) : inj(A,C)";
+Goal "[| g: inj(A,B);  f: inj(B,C) |] ==> (f O g) : inj(A,C)";
 by (res_inst_tac [("d", "%y. converse(g) ` (converse(f) ` y)")]
     f_imp_injective 1);
 by (REPEAT (ares_tac [comp_fun, inj_is_fun] 1));
@@ -403,7 +403,7 @@
 
 (*left inverse of composition; one inclusion is
         f: A->B ==> id(A) <= converse(f) O f *)
-Goalw [inj_def] "!!f. f: inj(A,B) ==> converse(f) O f = id(A)";
+Goalw [inj_def] "f: inj(A,B) ==> converse(f) O f = id(A)";
 by (fast_tac (claset() addIs [apply_Pair] 
                       addEs [domain_type]
                addss (simpset() addsimps [apply_iff])) 1);
@@ -442,11 +442,11 @@
        ORELSE eresolve_tac [bspec, apply_type] 1));
 qed "fg_imp_bijective";
 
-Goal "!!f A. [| f: A->A;  f O f = id(A) |] ==> f : bij(A,A)";
+Goal "[| f: A->A;  f O f = id(A) |] ==> f : bij(A,A)";
 by (REPEAT (ares_tac [fg_imp_bijective] 1));
 qed "nilpotent_imp_bijective";
 
-Goal "!!f A B. [| converse(f): B->A;  f: A->B |] ==> f : bij(A,B)";
+Goal "[| converse(f): B->A;  f: A->B |] ==> f : bij(A,B)";
 by (asm_simp_tac (simpset() addsimps [fg_imp_bijective, comp_eq_id_iff, 
                                   left_inverse_lemma, right_inverse_lemma]) 1);
 qed "invertible_imp_bijective";
@@ -491,7 +491,7 @@
 by (fast_tac (claset() addIs rls) 1);
 qed "surj_image";
 
-Goal "!!f. [| f: Pi(C,B);  A<=C |] ==> restrict(f,A)``A = f``A";
+Goal "[| f: Pi(C,B);  A<=C |] ==> restrict(f,A)``A = f``A";
 by (rtac equalityI 1);
 by (SELECT_GOAL (rewtac restrict_def) 2);
 by (REPEAT (eresolve_tac [imageE, apply_equality RS subst] 2
@@ -522,7 +522,7 @@
 
 (*** Lemmas for Ramsey's Theorem ***)
 
-Goalw [inj_def] "!!f. [| f: inj(A,B);  B<=D |] ==> f: inj(A,D)";
+Goalw [inj_def] "[| f: inj(A,B);  B<=D |] ==> f: inj(A,D)";
 by (blast_tac (claset() addIs [fun_weaken_type]) 1);
 qed "inj_weaken_type";
 
--- a/src/ZF/QPair.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/QPair.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -116,7 +116,7 @@
     "!!p. p:QSigma(A,B) ==> qsnd(p) : B(qfst(p))"
  (fn _=> [ Auto_tac ]);
 
-Goal "!!a A B. a: QSigma(A,B) ==> <qfst(a); qsnd(a)> = a";
+Goal "a: QSigma(A,B) ==> <qfst(a); qsnd(a)> = a";
 by Auto_tac;
 qed "QPair_qfst_qsnd_eq";
 
@@ -148,7 +148,7 @@
 
 (*** qsplit for predicates: result type o ***)
 
-Goalw [qsplit_def] "!!R a b. R(a,b) ==> qsplit(R, <a;b>)";
+Goalw [qsplit_def] "R(a,b) ==> qsplit(R, <a;b>)";
 by (Asm_simp_tac 1);
 qed "qsplitI";
 
@@ -162,7 +162,7 @@
 by (Asm_full_simp_tac 1);
 qed "qsplitE";
 
-Goalw [qsplit_def] "!!R a b. qsplit(R,<a;b>) ==> R(a,b)";
+Goalw [qsplit_def] "qsplit(R,<a;b>) ==> R(a,b)";
 by (Asm_full_simp_tac 1);
 qed "qsplitD";
 
@@ -211,11 +211,11 @@
 
 (** Introduction rules for the injections **)
 
-Goalw qsum_defs "!!a A B. a : A ==> QInl(a) : A <+> B";
+Goalw qsum_defs "a : A ==> QInl(a) : A <+> B";
 by (Blast_tac 1);
 qed "QInlI";
 
-Goalw qsum_defs "!!b A B. b : B ==> QInr(b) : A <+> B";
+Goalw qsum_defs "b : B ==> QInr(b) : A <+> B";
 by (Blast_tac 1);
 qed "QInrI";
 
@@ -266,11 +266,11 @@
 AddSDs [QInl_inject, QInr_inject];
 Addsimps [QInl_iff, QInr_iff, QInl_QInr_iff, QInr_QInl_iff, qsum_empty];
 
-Goal "!!A B. QInl(a): A<+>B ==> a: A";
+Goal "QInl(a): A<+>B ==> a: A";
 by (Blast_tac 1);
 qed "QInlD";
 
-Goal "!!A B. QInr(b): A<+>B ==> b: B";
+Goal "QInr(b): A<+>B ==> b: B";
 by (Blast_tac 1);
 qed "QInrD";
 
@@ -326,6 +326,6 @@
 by (Blast_tac 1);
 qed "Part_QInr2";
 
-Goal "!!A B C. C <= A <+> B ==> Part(C,QInl) Un Part(C,QInr) = C";
+Goal "C <= A <+> B ==> Part(C,QInl) Un Part(C,QInr) = C";
 by (Blast_tac 1);
 qed "Part_qsum_equality";
--- a/src/ZF/QUniv.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/QUniv.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -35,7 +35,7 @@
 by (etac PowD 1);
 qed "qunivD";
 
-Goalw [quniv_def] "!!A B. A<=B ==> quniv(A) <= quniv(B)";
+Goalw [quniv_def] "A<=B ==> quniv(A) <= quniv(B)";
 by (etac (eclose_mono RS univ_mono RS Pow_mono) 1);
 qed "quniv_mono";
 
@@ -80,7 +80,7 @@
 
 (** Quine disjoint sum **)
 
-Goalw [QInl_def] "!!A a. a <= univ(A) ==> QInl(a) <= univ(A)";
+Goalw [QInl_def] "a <= univ(A) ==> QInl(a) <= univ(A)";
 by (etac (empty_subsetI RS QPair_subset_univ) 1);
 qed "QInl_subset_univ";
 
@@ -91,7 +91,7 @@
 val naturals_subset_univ = 
     [naturals_subset_nat, nat_subset_univ] MRS subset_trans;
 
-Goalw [QInr_def] "!!A a. a <= univ(A) ==> QInr(a) <= univ(A)";
+Goalw [QInr_def] "a <= univ(A) ==> QInr(a) <= univ(A)";
 by (etac (nat_1I RS naturals_subset_univ RS QPair_subset_univ) 1);
 qed "QInr_subset_univ";
 
@@ -130,11 +130,11 @@
 
 (** Quine disjoint sum **)
 
-Goalw [QInl_def] "!!A a. a: quniv(A) ==> QInl(a) : quniv(A)";
+Goalw [QInl_def] "a: quniv(A) ==> QInl(a) : quniv(A)";
 by (REPEAT (ares_tac [zero_in_quniv,QPair_in_quniv] 1));
 qed "QInl_in_quniv";
 
-Goalw [QInr_def] "!!A b. b: quniv(A) ==> QInr(b) : quniv(A)";
+Goalw [QInr_def] "b: quniv(A) ==> QInr(b) : quniv(A)";
 by (REPEAT (ares_tac [one_in_quniv, QPair_in_quniv] 1));
 qed "QInr_in_quniv";
 
--- a/src/ZF/Rel.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Rel.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -29,7 +29,7 @@
 by (REPEAT (ares_tac (prems @ [allI,impI]) 1));
 qed "symI";
 
-Goalw [sym_def] "!!r. [| sym(r); <x,y>: r |]  ==>  <y,x>: r";
+Goalw [sym_def] "[| sym(r); <x,y>: r |]  ==>  <y,x>: r";
 by (Blast_tac 1);
 qed "symE";
 
--- a/src/ZF/Resid/Reduction.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Resid/Reduction.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -37,7 +37,7 @@
 (*     Lemmas for reduction                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-Goal  "!!u. m--->n ==> Fun(m) ---> Fun(n)";
+Goal  "m--->n ==> Fun(m) ---> Fun(n)";
 by (etac Sred.induct 1);
 by (resolve_tac [Sred.trans] 3);
 by (ALLGOALS (Asm_simp_tac ));
@@ -76,23 +76,23 @@
 (* ------------------------------------------------------------------------- *)
 
 
-Goal "!!u. m:lambda==> m =1=> m";
+Goal "m:lambda==> m =1=> m";
 by (eresolve_tac [lambda.induct] 1);
 by (ALLGOALS (Asm_simp_tac ));
 qed "refl_par_red1";
 
-Goal "!!u. m-1->n ==> m=1=>n";
+Goal "m-1->n ==> m=1=>n";
 by (etac Sred1.induct 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [refl_par_red1]) ));
 qed "red1_par_red1";
 
-Goal "!!u. m--->n ==> m===>n";
+Goal "m--->n ==> m===>n";
 by (etac Sred.induct 1);
 by (resolve_tac [Spar_red.trans] 3);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [refl_par_red1,red1_par_red1]) ));
 qed "red_par_red";
 
-Goal "!!u. m===>n ==> m--->n";
+Goal "m===>n ==> m--->n";
 by (etac Spar_red.induct 1);
 by (etac Spar_red1.induct 1);
 by (resolve_tac [Sred.trans] 5);
--- a/src/ZF/Resid/SubUnion.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Resid/SubUnion.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -65,7 +65,7 @@
 (*    comp proofs                                                            *)
 (* ------------------------------------------------------------------------- *)
 
-Goal "!!u. u:redexes ==> u ~ u";
+Goal "u:redexes ==> u ~ u";
 by (eresolve_tac [redexes.induct] 1);
 by (ALLGOALS Fast_tac);
 qed "comp_refl";
--- a/src/ZF/Sum.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Sum.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -48,11 +48,11 @@
 
 (** Introduction rules for the injections **)
 
-Goalw sum_defs "!!a A B. a : A ==> Inl(a) : A+B";
+Goalw sum_defs "a : A ==> Inl(a) : A+B";
 by (Blast_tac 1);
 qed "InlI";
 
-Goalw sum_defs "!!b A B. b : B ==> Inr(b) : A+B";
+Goalw sum_defs "b : B ==> Inr(b) : A+B";
 by (Blast_tac 1);
 qed "InrI";
 
@@ -103,11 +103,11 @@
 
 Addsimps [InlI, InrI, Inl_iff, Inr_iff, Inl_Inr_iff, Inr_Inl_iff, sum_empty];
 
-Goal "!!A B. Inl(a): A+B ==> a: A";
+Goal "Inl(a): A+B ==> a: A";
 by (Blast_tac 1);
 qed "InlD";
 
-Goal "!!A B. Inr(b): A+B ==> b: B";
+Goal "Inr(b): A+B ==> b: B";
 by (Blast_tac 1);
 qed "InrD";
 
@@ -178,7 +178,7 @@
 
 (*** More rules for Part(A,h) ***)
 
-Goal "!!A B h. A<=B ==> Part(A,h)<=Part(B,h)";
+Goal "A<=B ==> Part(A,h)<=Part(B,h)";
 by (Blast_tac 1);
 qed "Part_mono";
 
@@ -196,7 +196,7 @@
 by (Blast_tac 1);
 qed "Part_Inr";
 
-Goalw [Part_def] "!!a. a : Part(A,h) ==> a : A";
+Goalw [Part_def] "a : Part(A,h) ==> a : A";
 by (etac CollectD1 1);
 qed "PartD1";
 
@@ -208,6 +208,6 @@
 by (Blast_tac 1);
 qed "Part_Inr2";
 
-Goal "!!A B C. C <= A+B ==> Part(C,Inl) Un Part(C,Inr) = C";
+Goal "C <= A+B ==> Part(C,Inl) Un Part(C,Inr) = C";
 by (Blast_tac 1);
 qed "Part_sum_equality";
--- a/src/ZF/Trancl.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Trancl.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -117,17 +117,17 @@
 
 (** Conversions between trancl and rtrancl **)
 
-Goalw [trancl_def] "!!r. <a,b> : r^+ ==> <a,b> : r^*";
+Goalw [trancl_def] "<a,b> : r^+ ==> <a,b> : r^*";
 by (blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1);
 qed "trancl_into_rtrancl";
 
 (*r^+ contains all pairs in r  *)
-Goalw [trancl_def] "!!r. <a,b> : r ==> <a,b> : r^+";
+Goalw [trancl_def] "<a,b> : r ==> <a,b> : r^+";
 by (blast_tac (claset() addSIs [rtrancl_refl]) 1);
 qed "r_into_trancl";
 
 (*The premise ensures that r consists entirely of pairs*)
-Goal "!!r. r <= Sigma(A,B) ==> r <= r^+";
+Goal "r <= Sigma(A,B) ==> r <= r^+";
 by (blast_tac (claset() addIs [r_into_trancl]) 1);
 qed "r_subset_trancl";
 
--- a/src/ZF/Univ.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Univ.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -16,7 +16,7 @@
 
 (** Monotonicity **)
 
-Goal "!!A B. A<=B ==> ALL j. i<=j --> Vfrom(A,i) <= Vfrom(B,j)";
+Goal "A<=B ==> ALL j. i<=j --> Vfrom(A,i) <= Vfrom(B,j)";
 by (eps_ind_tac "i" 1);
 by (rtac (impI RS allI) 1);
 by (stac Vfrom 1);
@@ -66,7 +66,7 @@
 
 (*** Basic closure properties ***)
 
-Goal "!!x y. y:x ==> 0 : Vfrom(A,x)";
+Goal "y:x ==> 0 : Vfrom(A,x)";
 by (stac Vfrom 1);
 by (Blast_tac 1);
 qed "zero_in_Vfrom";
@@ -84,14 +84,14 @@
 
 bind_thm ("A_into_Vfrom", A_subset_Vfrom RS subsetD);
 
-Goal "!!A a i. a <= Vfrom(A,i) ==> a: Vfrom(A,succ(i))";
+Goal "a <= Vfrom(A,i) ==> a: Vfrom(A,succ(i))";
 by (stac Vfrom 1);
 by (Blast_tac 1);
 qed "subset_mem_Vfrom";
 
 (** Finite sets and ordered pairs **)
 
-Goal "!!a. a: Vfrom(A,i) ==> {a} : Vfrom(A,succ(i))";
+Goal "a: Vfrom(A,i) ==> {a} : Vfrom(A,succ(i))";
 by (rtac subset_mem_Vfrom 1);
 by Safe_tac;
 qed "singleton_in_Vfrom";
@@ -122,7 +122,7 @@
 by (Blast_tac 1);
 qed "Vfrom_0";
 
-Goal "!!i. Ord(i) ==> Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))";
+Goal "Ord(i) ==> Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))";
 by (rtac (Vfrom RS trans) 1);
 by (rtac (succI1 RS RepFunI RS Union_upper RSN
               (2, equalityI RS subst_context)) 1);
@@ -171,7 +171,7 @@
 by (rtac refl 1);
 qed "Limit_Vfrom_eq";
 
-Goal "!!a. [| a: Vfrom(A,j);  Limit(i);  j<i |] ==> a : Vfrom(A,i)";
+Goal "[| a: Vfrom(A,j);  Limit(i);  j<i |] ==> a : Vfrom(A,i)";
 by (rtac (Limit_Vfrom_eq RS equalityD2 RS subsetD) 1);
 by (REPEAT (ares_tac [ltD RS UN_I] 1));
 qed "Limit_VfromI";
@@ -221,7 +221,7 @@
 by (REPEAT (ares_tac [limiti, Limit_has_succ, Un_least_lt] 1));
 qed "Pair_in_VLimit";
 
-Goal "!!i. Limit(i) ==> Vfrom(A,i)*Vfrom(A,i) <= Vfrom(A,i)";
+Goal "Limit(i) ==> Vfrom(A,i)*Vfrom(A,i) <= Vfrom(A,i)";
 by (REPEAT (ares_tac [subsetI,Pair_in_VLimit] 1
      ORELSE eresolve_tac [SigmaE, ssubst] 1));
 qed "product_VLimit";
@@ -232,7 +232,7 @@
 bind_thm ("nat_subset_VLimit", 
           [nat_le_Limit RS le_imp_subset, i_subset_Vfrom] MRS subset_trans);
 
-Goal "!!i. [| n: nat;  Limit(i) |] ==> n : Vfrom(A,i)";
+Goal "[| n: nat;  Limit(i) |] ==> n : Vfrom(A,i)";
 by (REPEAT (ares_tac [nat_subset_VLimit RS subsetD] 1));
 qed "nat_into_VLimit";
 
@@ -240,7 +240,7 @@
 
 bind_thm ("zero_in_VLimit", Limit_has_0 RS ltD RS zero_in_Vfrom);
 
-Goal "!!i. Limit(i) ==> 1 : Vfrom(A,i)";
+Goal "Limit(i) ==> 1 : Vfrom(A,i)";
 by (REPEAT (ares_tac [nat_into_VLimit, nat_0I, nat_succI] 1));
 qed "one_in_VLimit";
 
@@ -254,7 +254,7 @@
 by (REPEAT (ares_tac [one_in_VLimit, Pair_in_VLimit] 1));
 qed "Inr_in_VLimit";
 
-Goal "!!i. Limit(i) ==> Vfrom(C,i)+Vfrom(C,i) <= Vfrom(C,i)";
+Goal "Limit(i) ==> Vfrom(C,i)+Vfrom(C,i) <= Vfrom(C,i)";
 by (blast_tac (claset() addSIs [Inl_in_VLimit, Inr_in_VLimit]) 1);
 qed "sum_VLimit";
 
@@ -264,14 +264,14 @@
 
 (*** Properties assuming Transset(A) ***)
 
-Goal "!!i A. Transset(A) ==> Transset(Vfrom(A,i))";
+Goal "Transset(A) ==> Transset(Vfrom(A,i))";
 by (eps_ind_tac "i" 1);
 by (stac Vfrom 1);
 by (blast_tac (claset() addSIs [Transset_Union_family, Transset_Un,
                             Transset_Pow]) 1);
 qed "Transset_Vfrom";
 
-Goal "!!A i. Transset(A) ==> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))";
+Goal "Transset(A) ==> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))";
 by (rtac (Vfrom_succ RS trans) 1);
 by (rtac (Un_upper2 RSN (2,equalityI)) 1);
 by (rtac (subset_refl RSN (2,Un_least)) 1);
@@ -427,13 +427,13 @@
 by (blast_tac (claset() addSIs [rank_lt RS ltD]) 1);
 val lemma = result();
 
-Goal "!!x i. rank(x)<i ==> x : Vset(i)";
+Goal "rank(x)<i ==> x : Vset(i)";
 by (etac ltE 1);
 by (etac (lemma RS spec RS mp) 1);
 by (assume_tac 1);
 qed "VsetI";
 
-Goal "!!i. Ord(i) ==> b : Vset(i) <-> rank(b) < i";
+Goal "Ord(i) ==> b : Vset(i) <-> rank(b) < i";
 by (rtac iffI 1);
 by (REPEAT (eresolve_tac [asm_rl, VsetD, VsetI] 1));
 qed "Vset_Ord_rank_iff";
@@ -443,7 +443,7 @@
 by (rtac (Ord_rank RS Vset_Ord_rank_iff) 1);
 qed "Vset_rank_iff";
 
-Goal "!!i. Ord(i) ==> rank(Vset(i)) = i";
+Goal "Ord(i) ==> rank(Vset(i)) = i";
 by (stac rank 1);
 by (rtac equalityI 1);
 by Safe_tac;
@@ -504,12 +504,12 @@
 
 (*** univ(A) ***)
 
-Goalw [univ_def] "!!A B. A<=B ==> univ(A) <= univ(B)";
+Goalw [univ_def] "A<=B ==> univ(A) <= univ(B)";
 by (etac Vfrom_mono 1);
 by (rtac subset_refl 1);
 qed "univ_mono";
 
-Goalw [univ_def] "!!A. Transset(A) ==> Transset(univ(A))";
+Goalw [univ_def] "Transset(A) ==> Transset(univ(A))";
 by (etac Transset_Vfrom 1);
 qed "Transset_univ";
 
@@ -519,7 +519,7 @@
 by (rtac (Limit_nat RS Limit_Vfrom_eq) 1);
 qed "univ_eq_UN";
 
-Goal "!!c. c <= univ(A) ==> c = (UN i:nat. c Int Vfrom(A,i))";
+Goal "c <= univ(A) ==> c = (UN i:nat. c Int Vfrom(A,i))";
 by (rtac (subset_UN_iff_eq RS iffD1) 1);
 by (etac (univ_eq_UN RS subst) 1);
 qed "subset_univ_eq_Int";
@@ -558,7 +558,7 @@
 
 (** Closure under unordered and ordered pairs **)
 
-Goalw [univ_def] "!!A a. a: univ(A) ==> {a} : univ(A)";
+Goalw [univ_def] "a: univ(A) ==> {a} : univ(A)";
 by (REPEAT (ares_tac [singleton_in_VLimit, Limit_nat] 1));
 qed "singleton_in_univ";
 
@@ -606,11 +606,11 @@
 
 (** Closure under disjoint union **)
 
-Goalw [univ_def] "!!A a. a: univ(A) ==> Inl(a) : univ(A)";
+Goalw [univ_def] "a: univ(A) ==> Inl(a) : univ(A)";
 by (etac (Limit_nat RSN (2,Inl_in_VLimit)) 1);
 qed "Inl_in_univ";
 
-Goalw [univ_def] "!!A b. b: univ(A) ==> Inr(b) : univ(A)";
+Goalw [univ_def] "b: univ(A) ==> Inr(b) : univ(A)";
 by (etac (Limit_nat RSN (2,Inr_in_VLimit)) 1);
 qed "Inr_in_univ";
 
@@ -640,7 +640,7 @@
 by (blast_tac (claset() addSIs [Un_least_lt] addIs [Vfrom_UnI1, Vfrom_UnI2]) 1);
 val Fin_Vfrom_lemma = result();
 
-Goal "!!i. Limit(i) ==> Fin(Vfrom(A,i)) <= Vfrom(A,i)";
+Goal "Limit(i) ==> Fin(Vfrom(A,i)) <= Vfrom(A,i)";
 by (rtac subsetI 1);
 by (dtac Fin_Vfrom_lemma 1);
 by Safe_tac;
@@ -665,7 +665,7 @@
 
 bind_thm ("nat_fun_subset_VLimit", [Pi_mono, nat_fun_VLimit] MRS subset_trans);
 
-Goalw [univ_def] "!!i. n: nat ==> n -> univ(A) <= univ(A)";
+Goalw [univ_def] "n: nat ==> n -> univ(A) <= univ(A)";
 by (etac (Limit_nat RSN (2,nat_fun_VLimit)) 1);
 val nat_fun_univ = result();
 
--- a/src/ZF/WF.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/WF.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -136,24 +136,24 @@
 
 (*** Properties of well-founded relations ***)
 
-Goal "!!r. wf(r) ==> <a,a> ~: r";
+Goal "wf(r) ==> <a,a> ~: r";
 by (wf_ind_tac "a" [] 1);
 by (Blast_tac 1);
 qed "wf_not_refl";
 
-Goal "!!r. [| wf(r);  <a,x>:r;  <x,a>:r |] ==> P";
+Goal "[| wf(r);  <a,x>:r;  <x,a>:r |] ==> P";
 by (subgoal_tac "ALL x. <a,x>:r --> <x,a>:r --> P" 1);
 by (wf_ind_tac "a" [] 2);
 by (Blast_tac 2);
 by (Blast_tac 1);
 qed "wf_asym";
 
-Goal "!!r. [| wf[A](r); a: A |] ==> <a,a> ~: r";
+Goal "[| wf[A](r); a: A |] ==> <a,a> ~: r";
 by (wf_on_ind_tac "a" [] 1);
 by (Blast_tac 1);
 qed "wf_on_not_refl";
 
-Goal "!!r. [| wf[A](r);  <a,b>:r;  <b,a>:r;  a:A;  b:A |] ==> P";
+Goal "[| wf[A](r);  <a,b>:r;  <b,a>:r;  a:A;  b:A |] ==> P";
 by (subgoal_tac "ALL y:A. <a,y>:r --> <y,a>:r --> P" 1);
 by (wf_on_ind_tac "a" [] 2);
 by (Blast_tac 2);
@@ -185,7 +185,7 @@
 by (blast_tac (claset() addEs [tranclE]) 1);
 qed "wf_on_trancl";
 
-Goal "!!r. wf(r) ==> wf(r^+)";
+Goal "wf(r) ==> wf(r^+)";
 by (asm_full_simp_tac (simpset() addsimps [wf_iff_wf_on_field]) 1);
 by (rtac (trancl_type RS field_rel_subset RSN (2, wf_on_subset_A)) 1);
 by (etac wf_on_trancl 1);
--- a/src/ZF/ZF.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ZF.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -9,7 +9,7 @@
 open ZF;
 
 (*Useful examples:  singletonI RS subst_elem,  subst_elem RSN (2,IntI) *)
-Goal "!!a b A. [| b:A;  a=b |] ==> a:A";
+Goal "[| b:A;  a=b |] ==> a:A";
 by (etac ssubst 1);
 by (assume_tac 1);
 val subst_elem = result();
--- a/src/ZF/Zorn.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/Zorn.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -199,13 +199,12 @@
 by (rewtac increasing_def);
 by (rtac CollectI 1);
 by (rtac lam_type 1);
-by (asm_simp_tac (simpset() addsplits [split_if]) 1);
+by (Asm_simp_tac 1);
 by (fast_tac (claset() addSIs [super_subset_chain RS subsetD,
 			      chain_subset_Pow RS subsetD,
 			      choice_super]) 1);
 (*Now, verify that it increases*)
-by (asm_simp_tac (simpset() addsimps [Pow_iff, subset_refl]
-                        addsplits [split_if]) 1);
+by (asm_simp_tac (simpset() addsimps [Pow_iff, subset_refl]) 1);
 by Safe_tac;
 by (dtac choice_super 1);
 by (REPEAT (assume_tac 1));
@@ -224,8 +223,7 @@
 by (etac TFin_induct 1);
 by (asm_simp_tac 
     (simpset() addsimps [chain_subset_Pow RS subsetD, 
-                     choice_super RS (super_subset_chain RS subsetD)]
-           addsplits [split_if]) 1);
+			 choice_super RS (super_subset_chain RS subsetD)]) 1);
 by (rewtac chain_def);
 by (rtac CollectI 1 THEN Blast_tac 1);
 by Safe_tac;
@@ -249,8 +247,7 @@
 by (rtac refl 2);
 by (asm_full_simp_tac 
     (simpset() addsimps [subset_refl RS TFin_UnionI RS
-                     (TFin.dom_subset RS subsetD)]
-           addsplits [split_if]) 1);
+			 (TFin.dom_subset RS subsetD)]) 1);
 by (eresolve_tac [choice_not_equals RS notE] 1);
 by (REPEAT (assume_tac 1));
 qed "Hausdorff";
@@ -302,7 +299,7 @@
 qed "TFin_well_lemma5";
 
 (*Well-ordering of TFin(S,next)*)
-Goal "!!Z. [| Z <= TFin(S,next);  z:Z |] ==> Inter(Z) : Z";
+Goal "[| Z <= TFin(S,next);  z:Z |] ==> Inter(Z) : Z";
 by (rtac classical 1);
 by (subgoal_tac "Z = {Union(TFin(S,next))}" 1);
 by (asm_simp_tac (simpset() addsimps [Inter_singleton]) 1);
@@ -358,11 +355,10 @@
 (*Verify that it increases*)
 by (rtac allI 2);
 by (rtac impI 2);
-by (asm_simp_tac (simpset() addsimps [Pow_iff, subset_consI, subset_refl]
-                        addsplits [split_if]) 2);
+by (asm_simp_tac (simpset() addsimps [Pow_iff, subset_consI, subset_refl]) 2);
 (*Type checking is surprisingly hard!*)
-by (asm_simp_tac (simpset() addsimps [Pow_iff, cons_subset_iff, subset_refl]
-                        addsplits [split_if]) 1);
+by (asm_simp_tac
+    (simpset() addsimps [Pow_iff, cons_subset_iff, subset_refl]) 1);
 by (blast_tac (claset() addSIs [choice_Diff RS DiffD1]) 1);
 qed "Zermelo_next_exists";
 
@@ -391,15 +387,13 @@
     (simpset() delsimps [Union_iff]
               addsimps [Collect_subset RS TFin_UnionI RS TFin_is_subset,
                      Pow_iff, cons_subset_iff, subset_refl,
-                     choice_Diff RS DiffD2]
-           addsplits [split_if]) 2);
+                     choice_Diff RS DiffD2]) 2);
 by (subgoal_tac "x : next ` Union({y: TFin(S,next). x~: y})" 1);
 by (blast_tac (subset_cs addSIs [Collect_subset RS TFin_UnionI, TFin_nextI]) 2);
 (*End of the lemmas!*)
 by (asm_full_simp_tac 
     (simpset() addsimps [Collect_subset RS TFin_UnionI RS TFin_is_subset,
-                     Pow_iff, cons_subset_iff, subset_refl]
-           addsplits [split_if]) 1);
+                     Pow_iff, cons_subset_iff, subset_refl]) 1);
 by (REPEAT (eresolve_tac [asm_rl, consE, sym, notE] 1));
 qed "choice_imp_injection";
 
--- a/src/ZF/ex/Acc.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/Acc.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -20,7 +20,7 @@
 by (fast_tac (claset() addIs (prems@acc.intrs)) 1);
 qed "accI";
 
-Goal "!!a b r. [| b: acc(r);  <a,b>: r |] ==> a: acc(r)";
+Goal "[| b: acc(r);  <a,b>: r |] ==> a: acc(r)";
 by (etac acc.elim 1);
 by (Fast_tac 1);
 qed "acc_downward";
--- a/src/ZF/ex/BT.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/BT.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -19,7 +19,7 @@
 
 (**  Lemmas to justify using "bt" in other recursive type definitions **)
 
-Goalw bt.defs "!!A B. A<=B ==> bt(A) <= bt(B)";
+Goalw bt.defs "A<=B ==> bt(A) <= bt(B)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac bt.bnd_mono 1));
 by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
@@ -110,7 +110,7 @@
 val [bt_reflect_Lf, bt_reflect_Br] = bt_recs bt_reflect_def;
 Addsimps [bt_reflect_Lf, bt_reflect_Br];
 
-Goalw [bt_reflect_def] "!!xs. xs: bt(A) ==> bt_reflect(xs) : bt(A)";
+Goalw [bt_reflect_def] "xs: bt(A) ==> bt_reflect(xs) : bt(A)";
 by (REPEAT (ares_tac (bt.intrs @ [bt_rec_type]) 1));
 qed "bt_reflect_type";
 
--- a/src/ZF/ex/Bin.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/Bin.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -227,11 +227,11 @@
 
 (*** bin_add: binary addition ***)
 
-Goalw [bin_add_def] "!!w. w: bin ==> bin_add(Plus,w) = w";
+Goalw [bin_add_def] "w: bin ==> bin_add(Plus,w) = w";
 by (Asm_simp_tac 1);
 qed "bin_add_Plus";
 
-Goalw [bin_add_def] "!!w. w: bin ==> bin_add(Minus,w) = bin_pred(w)";
+Goalw [bin_add_def] "w: bin ==> bin_add(Minus,w) = bin_pred(w)";
 by (Asm_simp_tac 1);
 qed "bin_add_Minus";
 
--- a/src/ZF/ex/CoUnit.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/CoUnit.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -66,7 +66,7 @@
 val counit2_Int_Vset_subset = standard
         (counit2_Int_Vset_subset_lemma RS spec RS spec RS mp);
 
-Goal "!!x y. [| x: counit2;  y: counit2 |] ==> x=y";
+Goal "[| x: counit2;  y: counit2 |] ==> x=y";
 by (rtac equalityI 1);
 by (REPEAT (ares_tac [conjI, counit2_Int_Vset_subset RS Int_Vset_subset] 1));
 qed "counit2_implies_equal";
--- a/src/ZF/ex/Comb.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/Comb.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -70,7 +70,7 @@
                      contract.Ap2 RS rtrancl_into_rtrancl2];
 
 (*Example only: not used*)
-Goalw [I_def] "!!p. p:comb ==> I#p ---> p";
+Goalw [I_def] "p:comb ==> I#p ---> p";
 by (blast_tac (claset() addIs reduction_rls) 1);
 qed "reduce_I";
 
@@ -88,15 +88,15 @@
 AddIs  [contract.Ap1, contract.Ap2];
 AddSEs [K_contractE, S_contractE, Ap_contractE];
 
-Goalw [I_def] "!!r. I -1-> r ==> P";
+Goalw [I_def] "I -1-> r ==> P";
 by (Blast_tac 1);
 qed "I_contract_E";
 
-Goal "!!p r. K#p -1-> r ==> (EX q. r = K#q & p -1-> q)";
+Goal "K#p -1-> r ==> (EX q. r = K#q & p -1-> q)";
 by (Blast_tac 1);
 qed "K1_contractD";
 
-Goal "!!p r. [| p ---> q;  r: comb |] ==> p#r ---> q#r";
+Goal "[| p ---> q;  r: comb |] ==> p#r ---> q#r";
 by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1);
 by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1);
 by (etac rtrancl_induct 1);
@@ -105,7 +105,7 @@
 by (blast_tac (claset() addIs (contract_combD2::reduction_rls)) 1);
 qed "Ap_reduce1";
 
-Goal "!!p r. [| p ---> q;  r: comb |] ==> r#p ---> r#q";
+Goal "[| p ---> q;  r: comb |] ==> r#p ---> r#q";
 by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1);
 by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1);
 by (etac rtrancl_induct 1);
@@ -157,11 +157,11 @@
 
 (*** Basic properties of parallel contraction ***)
 
-Goal "!!p r. K#p =1=> r ==> (EX p'. r = K#p' & p =1=> p')";
+Goal "K#p =1=> r ==> (EX p'. r = K#p' & p =1=> p')";
 by (Blast_tac 1);
 qed "K1_parcontractD";
 
-Goal "!!p r. S#p =1=> r ==> (EX p'. r = S#p' & p =1=> p')";
+Goal "S#p =1=> r ==> (EX p'. r = S#p' & p =1=> p')";
 by (Blast_tac 1);
 qed "S1_parcontractD";
 
@@ -181,12 +181,12 @@
 
 (*** Equivalence of p--->q and p===>q ***)
 
-Goal "!!p q. p-1->q ==> p=1=>q";
+Goal "p-1->q ==> p=1=>q";
 by (etac contract.induct 1);
 by (ALLGOALS Blast_tac);
 qed "contract_imp_parcontract";
 
-Goal "!!p q. p--->q ==> p===>q";
+Goal "p--->q ==> p===>q";
 by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1);
 by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1);
 by (etac rtrancl_induct 1);
@@ -196,7 +196,7 @@
 qed "reduce_imp_parreduce";
 
 
-Goal "!!p q. p=1=>q ==> p--->q";
+Goal "p=1=>q ==> p--->q";
 by (etac parcontract.induct 1);
 by (blast_tac (claset() addIs reduction_rls) 1);
 by (blast_tac (claset() addIs reduction_rls) 1);
@@ -207,7 +207,7 @@
 		    parcontract_combD2]) 1);
 qed "parcontract_imp_reduce";
 
-Goal "!!p q. p===>q ==> p--->q";
+Goal "p===>q ==> p--->q";
 by (forward_tac [trancl_type RS subsetD RS SigmaD1] 1);
 by (dtac (field_parcontract_eq RS equalityD1 RS subsetD) 1);
 by (etac trancl_induct 1);
--- a/src/ZF/ex/Data.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/Data.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -17,7 +17,7 @@
 
 (**  Lemmas to justify using "data" in other recursive type definitions **)
 
-Goalw data.defs "!!A B. [| A<=C; B<=D |] ==> data(A,B) <= data(C,D)";
+Goalw data.defs "[| A<=C; B<=D |] ==> data(A,B) <= data(C,D)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac data.bnd_mono 1));
 by (REPEAT (ares_tac (univ_mono::Un_mono::basic_monos) 1));
--- a/src/ZF/ex/Integ.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/Integ.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -38,7 +38,7 @@
 qed "intrel_iff";
 
 Goalw [intrel_def]
-    "!!x1 x2. [| x1#+y2 = x2#+y1; x1: nat; y1: nat; x2: nat; y2: nat |] ==> \
+    "[| x1#+y2 = x2#+y1; x1: nat; y1: nat; x2: nat; y2: nat |] ==> \
 \             <<x1,y1>,<x2,y2>>: intrel";
 by (fast_tac (claset() addIs prems) 1);
 qed "intrelI";
@@ -78,12 +78,11 @@
 (** znat: the injection from nat to integ **)
 
 Goalw [integ_def,quotient_def,znat_def]
-    "!!m. m : nat ==> $#m : integ";
+    "m : nat ==> $#m : integ";
 by (fast_tac (claset() addSIs [nat_0I]) 1);
 qed "znat_type";
 
-Goalw [znat_def]
-    "!!m n. [| $#m = $#n;  n: nat |] ==> m=n";
+Goalw [znat_def] "[| $#m = $#n;  n: nat |] ==> m=n";
 by (dtac eq_intrelD 1);
 by (typechk_tac [nat_0I, SigmaI]);
 by (Asm_full_simp_tac 1);
@@ -92,8 +91,7 @@
 
 (**** zminus: unary negation on integ ****)
 
-Goalw [congruent_def]
-    "congruent(intrel, %<x,y>. intrel``{<y,x>})";
+Goalw [congruent_def] "congruent(intrel, %<x,y>. intrel``{<y,x>})";
 by Safe_tac;
 by (asm_full_simp_tac (simpset() addsimps add_ac) 1);
 qed "zminus_congruent";
@@ -101,14 +99,12 @@
 (*Resolve th against the corresponding facts for zminus*)
 val zminus_ize = RSLIST [equiv_intrel, zminus_congruent];
 
-Goalw [integ_def,zminus_def]
-    "!!z. z : integ ==> $~z : integ";
+Goalw [integ_def,zminus_def] "z : integ ==> $~z : integ";
 by (typechk_tac [split_type, SigmaI, zminus_ize UN_equiv_class_type,
                  quotientI]);
 qed "zminus_type";
 
-Goalw [integ_def,zminus_def]
-    "!!z w. [| $~z = $~w;  z: integ;  w: integ |] ==> z=w";
+Goalw [integ_def,zminus_def] "[| $~z = $~w;  z: integ;  w: integ |] ==> z=w";
 by (etac (zminus_ize UN_equiv_class_inject) 1);
 by Safe_tac;
 (*The setloop is only needed because assumptions are in the wrong order!*)
@@ -117,11 +113,11 @@
 qed "zminus_inject";
 
 Goalw [zminus_def]
-    "!!x y.[| x: nat;  y: nat |] ==> $~ (intrel``{<x,y>}) = intrel `` {<y,x>}";
+    "[| x: nat;  y: nat |] ==> $~ (intrel``{<x,y>}) = intrel `` {<y,x>}";
 by (asm_simp_tac (simpset() addsimps [zminus_ize UN_equiv_class, SigmaI]) 1);
 qed "zminus";
 
-Goalw [integ_def] "!!z. z : integ ==> $~ ($~ z) = z";
+Goalw [integ_def] "z : integ ==> $~ ($~ z) = z";
 by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1));
 by (asm_simp_tac (simpset() addsimps [zminus]) 1);
 qed "zminus_zminus";
@@ -142,19 +138,15 @@
               (simpset() addsimps [add_le_self2 RS le_imp_not_lt])) 1);
 qed "not_znegative_znat";
 
-Goalw [znegative_def, znat_def]
-    "!!n. n: nat ==> znegative($~ $# succ(n))";
+Goalw [znegative_def, znat_def] "n: nat ==> znegative($~ $# succ(n))";
 by (asm_simp_tac (simpset() addsimps [zminus]) 1);
-by (REPEAT 
-    (ares_tac [refl, exI, conjI, nat_0_le,
-               refl RS intrelI RS imageI, consI1, nat_0I, nat_succI] 1));
+by(blast_tac (claset() addIs [nat_0_le]) 1);
 qed "znegative_zminus_znat";
 
 
 (**** zmagnitude: magnitide of an integer, as a natural number ****)
 
-Goalw [congruent_def]
-    "congruent(intrel, %<x,y>. (y#-x) #+ (x#-y))";
+Goalw [congruent_def] "congruent(intrel, %<x,y>. (y#-x) #+ (x#-y))";
 by Safe_tac;
 by (ALLGOALS Asm_simp_tac);
 by (etac rev_mp 1);
@@ -177,25 +169,23 @@
 val zmagnitude_ize = RSLIST [equiv_intrel, zmagnitude_congruent];
 
 Goalw [integ_def,zmagnitude_def]
-    "!!z. z : integ ==> zmagnitude(z) : nat";
+    "z : integ ==> zmagnitude(z) : nat";
 by (typechk_tac [split_type, zmagnitude_ize UN_equiv_class_type,
                  add_type, diff_type]);
 qed "zmagnitude_type";
 
 Goalw [zmagnitude_def]
-    "!!x y. [| x: nat;  y: nat |] ==> \
-\           zmagnitude (intrel``{<x,y>}) = (y #- x) #+ (x #- y)";
+    "[| x: nat;  y: nat |]  \
+\    ==> zmagnitude (intrel``{<x,y>}) = (y #- x) #+ (x #- y)";
 by (asm_simp_tac
     (simpset() addsimps [zmagnitude_ize UN_equiv_class, SigmaI]) 1);
 qed "zmagnitude";
 
-Goalw [znat_def]
-    "!!n. n: nat ==> zmagnitude($# n) = n";
+Goalw [znat_def] "n: nat ==> zmagnitude($# n) = n";
 by (asm_simp_tac (simpset() addsimps [zmagnitude]) 1);
 qed "zmagnitude_znat";
 
-Goalw [znat_def]
-    "!!n. n: nat ==> zmagnitude($~ $# n) = n";
+Goalw [znat_def] "n: nat ==> zmagnitude($~ $# n) = n";
 by (asm_simp_tac (simpset() addsimps [zmagnitude, zminus]) 1);
 qed "zmagnitude_zminus_znat";
 
@@ -222,71 +212,67 @@
 (*Resolve th against the corresponding facts for zadd*)
 val zadd_ize = RSLIST [equiv_intrel, zadd_congruent2];
 
-Goalw [integ_def,zadd_def]
-    "!!z w. [| z: integ;  w: integ |] ==> z $+ w : integ";
+Goalw [integ_def,zadd_def] "[| z: integ;  w: integ |] ==> z $+ w : integ";
 by (rtac (zadd_ize UN_equiv_class_type2) 1);
 by (simp_tac (simpset() addsimps [Let_def]) 3);
 by (REPEAT (ares_tac [split_type, add_type, quotientI, SigmaI] 1));
 qed "zadd_type";
 
 Goalw [zadd_def]
-  "!!x1 y1. [| x1: nat; y1: nat;  x2: nat; y2: nat |] ==>       \
+  "[| x1: nat; y1: nat;  x2: nat; y2: nat |] ==>       \
 \           (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) =        \
 \           intrel `` {<x1#+x2, y1#+y2>}";
 by (asm_simp_tac (simpset() addsimps [zadd_ize UN_equiv_class2, SigmaI]) 1);
 by (simp_tac (simpset() addsimps [Let_def]) 1);
 qed "zadd";
 
-Goalw [integ_def,znat_def] "!!z. z : integ ==> $#0 $+ z = z";
+Goalw [integ_def,znat_def] "z : integ ==> $#0 $+ z = z";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps [zadd]) 1);
 qed "zadd_0";
 
-Goalw [integ_def]
-    "!!z w. [| z: integ;  w: integ |] ==> $~ (z $+ w) = $~ z $+ $~ w";
+Goalw [integ_def] "[| z: integ;  w: integ |] ==> $~ (z $+ w) = $~ z $+ $~ w";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps [zminus,zadd]) 1);
 qed "zminus_zadd_distrib";
 
-Goalw [integ_def]
-    "!!z w. [| z: integ;  w: integ |] ==> z $+ w = w $+ z";
+Goalw [integ_def] "[| z: integ;  w: integ |] ==> z $+ w = w $+ z";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps (add_ac @ [zadd])) 1);
 qed "zadd_commute";
 
 Goalw [integ_def]
-    "!!z1 z2 z3. [| z1: integ;  z2: integ;  z3: integ |] ==> \
-\                (z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)";
+    "[| z1: integ;  z2: integ;  z3: integ |]   \
+\    ==> (z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 (*rewriting is much faster without intrel_iff, etc.*)
 by (asm_simp_tac (simpset() addsimps [zadd, add_assoc]) 1);
 qed "zadd_assoc";
 
 (*For AC rewriting*)
-qed_goal "zadd_left_commute" Integ.thy
-    "!!z1 z2 z3. [| z1:integ;  z2:integ;  z3: integ |] ==> \
-\                z1$+(z2$+z3) = z2$+(z1$+z3)"
- (fn _ => [asm_simp_tac (simpset() addsimps [zadd_assoc RS sym, zadd_commute]) 1]);
+Goal "[| z1:integ;  z2:integ;  z3: integ |] ==> z1$+(z2$+z3) = z2$+(z1$+z3)";
+by (asm_simp_tac (simpset() addsimps [zadd_assoc RS sym, zadd_commute]) 1);
+qed "zadd_left_commute";
 
 (*Integer addition is an AC operator*)
 val zadd_ac = [zadd_assoc, zadd_commute, zadd_left_commute];
 
 Goalw [znat_def]
-    "!!m n. [| m: nat;  n: nat |] ==> $# (m #+ n) = ($#m) $+ ($#n)";
+    "[| m: nat;  n: nat |] ==> $# (m #+ n) = ($#m) $+ ($#n)";
 by (asm_simp_tac (simpset() addsimps [zadd]) 1);
 qed "znat_add";
 
-Goalw [integ_def,znat_def] "!!z. z : integ ==> z $+ ($~ z) = $#0";
+Goalw [integ_def,znat_def] "z : integ ==> z $+ ($~ z) = $#0";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps [zminus, zadd, add_commute]) 1);
 qed "zadd_zminus_inverse";
 
-Goal "!!z. z : integ ==> ($~ z) $+ z = $#0";
+Goal "z : integ ==> ($~ z) $+ z = $#0";
 by (asm_simp_tac
     (simpset() addsimps [zadd_commute, zminus_type, zadd_zminus_inverse]) 1);
 qed "zadd_zminus_inverse2";
 
-Goal "!!z. z:integ ==> z $+ $#0 = z";
+Goal "z:integ ==> z $+ $#0 = z";
 by (rtac (zadd_commute RS trans) 1);
 by (REPEAT (ares_tac [znat_type, nat_0I, zadd_0] 1));
 qed "zadd_0_right";
@@ -320,52 +306,47 @@
 (*Resolve th against the corresponding facts for zmult*)
 val zmult_ize = RSLIST [equiv_intrel, zmult_congruent2];
 
-Goalw [integ_def,zmult_def]
-    "!!z w. [| z: integ;  w: integ |] ==> z $* w : integ";
+Goalw [integ_def,zmult_def] "[| z: integ;  w: integ |] ==> z $* w : integ";
 by (REPEAT (ares_tac [zmult_ize UN_equiv_class_type2,
                       split_type, add_type, mult_type, 
                       quotientI, SigmaI] 1));
 qed "zmult_type";
 
 Goalw [zmult_def]
-     "!!x1 x2. [| x1: nat; y1: nat;  x2: nat; y2: nat |] ==>    \
+     "[| x1: nat; y1: nat;  x2: nat; y2: nat |] ==>    \
 \              (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) =     \
 \              intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}";
 by (asm_simp_tac (simpset() addsimps [zmult_ize UN_equiv_class2, SigmaI]) 1);
 qed "zmult";
 
-Goalw [integ_def,znat_def] "!!z. z : integ ==> $#0 $* z = $#0";
+Goalw [integ_def,znat_def] "z : integ ==> $#0 $* z = $#0";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps [zmult]) 1);
 qed "zmult_0";
 
-Goalw [integ_def,znat_def]
-    "!!z. z : integ ==> $#1 $* z = z";
+Goalw [integ_def,znat_def] "z : integ ==> $#1 $* z = z";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps [zmult, add_0_right]) 1);
 qed "zmult_1";
 
-Goalw [integ_def]
-    "!!z w. [| z: integ;  w: integ |] ==> ($~ z) $* w = $~ (z $* w)";
+Goalw [integ_def] "[| z: integ;  w: integ |] ==> ($~ z) $* w = $~ (z $* w)";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps ([zminus, zmult] @ add_ac)) 1);
 qed "zmult_zminus";
 
-Goalw [integ_def]
-    "!!z w. [| z: integ;  w: integ |] ==> ($~ z) $* ($~ w) = (z $* w)";
+Goalw [integ_def] "[| z: integ;  w: integ |] ==> ($~ z) $* ($~ w) = (z $* w)";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps ([zminus, zmult] @ add_ac)) 1);
 qed "zmult_zminus_zminus";
 
-Goalw [integ_def]
-    "!!z w. [| z: integ;  w: integ |] ==> z $* w = w $* z";
+Goalw [integ_def] "[| z: integ;  w: integ |] ==> z $* w = w $* z";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps ([zmult] @ add_ac @ mult_ac)) 1);
 qed "zmult_commute";
 
 Goalw [integ_def]
-    "!!z1 z2 z3. [| z1: integ;  z2: integ;  z3: integ |] ==> \
-\                (z1 $* z2) $* z3 = z1 $* (z2 $* z3)";
+    "[| z1: integ;  z2: integ;  z3: integ |]     \
+\    ==> (z1 $* z2) $* z3 = z1 $* (z2 $* z3)";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac 
     (simpset() addsimps ([zmult, add_mult_distrib_left, 
@@ -373,17 +354,15 @@
 qed "zmult_assoc";
 
 (*For AC rewriting*)
-qed_goal "zmult_left_commute" Integ.thy
-    "!!z1 z2 z3. [| z1:integ;  z2:integ;  z3: integ |] ==> \
-\                z1$*(z2$*z3) = z2$*(z1$*z3)"
- (fn _ => [asm_simp_tac (simpset() addsimps [zmult_assoc RS sym, 
-                                         zmult_commute]) 1]);
+Goal "[| z1:integ;  z2:integ;  z3: integ |] ==> z1$*(z2$*z3) = z2$*(z1$*z3)";
+by (asm_simp_tac (simpset() addsimps [zmult_assoc RS sym, zmult_commute]) 1);
+qed "zmult_left_commute";
 
 (*Integer multiplication is an AC operator*)
 val zmult_ac = [zmult_assoc, zmult_commute, zmult_left_commute];
 
 Goalw [integ_def]
-    "!!z1 z2 z3. [| z1: integ;  z2: integ;  w: integ |] ==> \
+    "[| z1: integ;  z2: integ;  w: integ |] ==> \
 \                (z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps [zadd, zmult, add_mult_distrib]) 1);
--- a/src/ZF/ex/LList.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/LList.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -29,7 +29,7 @@
 
 (*** Lemmas to justify using "llist" in other recursive type definitions ***)
 
-Goalw llist.defs "!!A B. A<=B ==> llist(A) <= llist(B)";
+Goalw llist.defs "A<=B ==> llist(A) <= llist(B)";
 by (rtac gfp_mono 1);
 by (REPEAT (rtac llist.bnd_mono 1));
 by (REPEAT (ares_tac (quniv_mono::basic_monos) 1));
@@ -44,7 +44,7 @@
 AddSEs [Ord_in_Ord];
 
 Goal
-   "!!i. Ord(i) ==> ALL l: llist(quniv(A)). l Int Vset(i) <= univ(eclose(A))";
+   "Ord(i) ==> ALL l: llist(quniv(A)). l Int Vset(i) <= univ(eclose(A))";
 by (etac trans_induct 1);
 by (rtac ballI 1);
 by (etac llist.elim 1);
@@ -72,7 +72,7 @@
 
 (*Lemma for proving finality.  Unfold the lazy list; use induction hypothesis*)
 Goal
-   "!!i. Ord(i) ==> ALL l l'. <l,l'> : lleq(A) --> l Int Vset(i) <= l'";
+   "Ord(i) ==> ALL l l'. <l,l'> : lleq(A) --> l Int Vset(i) <= l'";
 by (etac trans_induct 1);
 by (REPEAT (resolve_tac [allI, impI] 1));
 by (etac lleq.elim 1);
@@ -94,7 +94,7 @@
 by (ALLGOALS Fast_tac);
 qed "lleq_symmetric";
 
-Goal "!!l l'. <l,l'> : lleq(A) ==> l=l'";
+Goal "<l,l'> : lleq(A) ==> l=l'";
 by (rtac equalityI 1);
 by (REPEAT (ares_tac [lleq_Int_Vset_subset RS Int_Vset_subset] 1
      ORELSE etac lleq_symmetric 1));
@@ -131,12 +131,12 @@
 
 bind_thm ("member_subset_Union_eclose", (arg_into_eclose RS Union_upper));
 
-Goal "!!a A. a : A ==> lconst(a) : quniv(A)";
+Goal "a : A ==> lconst(a) : quniv(A)";
 by (rtac (lconst_subset RS subset_trans RS qunivI) 1);
 by (etac (arg_into_eclose RS eclose_subset RS univ_mono) 1);
 qed "lconst_in_quniv";
 
-Goal "!!a A. a:A ==> lconst(a): llist(A)";
+Goal "a:A ==> lconst(a): llist(A)";
 by (rtac (singletonI RS llist.coinduct) 1);
 by (etac (lconst_in_quniv RS singleton_subsetI) 1);
 by (fast_tac (claset() addSIs [lconst]) 1);
@@ -156,7 +156,7 @@
 (*Reasoning borrowed from lleq.ML; a similar proof works for all
   "productive" functions -- cf Coquand's "Infinite Objects in Type Theory".*)
 Goal
-   "!!i. Ord(i) ==> ALL l: llist(bool). flip(l) Int Vset(i) <= \
+   "Ord(i) ==> ALL l: llist(bool). flip(l) Int Vset(i) <= \
 \                   univ(eclose(bool))";
 by (etac trans_induct 1);
 by (rtac ballI 1);
@@ -168,7 +168,7 @@
 by (deepen_tac (claset() addIs [Ord_trans, Int_lower1 RS subset_trans]) 2 1);
 qed "flip_llist_quniv_lemma";
 
-Goal "!!l. l: llist(bool) ==> flip(l) : quniv(bool)";
+Goal "l: llist(bool) ==> flip(l) : quniv(bool)";
 by (rtac (flip_llist_quniv_lemma RS bspec RS Int_Vset_subset RS qunivI) 1);
 by (REPEAT (assume_tac 1));
 qed "flip_in_quniv";
--- a/src/ZF/ex/ListN.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/ListN.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -11,7 +11,7 @@
 
 open ListN;
 
-Goal "!!l. l:list(A) ==> <length(l),l> : listn(A)";
+Goal "l:list(A) ==> <length(l),l> : listn(A)";
 by (etac list.induct 1);
 by (ALLGOALS Asm_simp_tac);
 by (REPEAT (ares_tac listn.intrs 1));
@@ -29,7 +29,7 @@
 by (simp_tac (simpset() addsimps [listn_iff,separation,image_singleton_iff]) 1);
 qed "listn_image_eq";
 
-Goalw listn.defs "!!A B. A<=B ==> listn(A) <= listn(B)";
+Goalw listn.defs "A<=B ==> listn(A) <= listn(B)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac listn.bnd_mono 1));
 by (REPEAT (ares_tac ([univ_mono,Sigma_mono,list_mono] @ basic_monos) 1));
--- a/src/ZF/ex/Mutil.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/Mutil.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -33,8 +33,7 @@
 Goalw [evnodd_def]
     "evnodd(cons(<i,j>,C), b) = \
 \    if((i#+j) mod 2 = b, cons(<i,j>, evnodd(C,b)), evnodd(C,b))";
-by (asm_simp_tac (simpset() addsimps [evnodd_def, Collect_cons] 
-                        addsplits [split_if]) 1);
+by (asm_simp_tac (simpset() addsimps [evnodd_def, Collect_cons]) 1);
 qed "evnodd_cons";
 
 Goalw [evnodd_def] "evnodd(0, b) = 0";
@@ -45,19 +44,18 @@
 
 (*** Dominoes ***)
 
-Goal "!!d. d:domino ==> Finite(d)";
+Goal "d:domino ==> Finite(d)";
 by (blast_tac (claset() addSIs [Finite_cons, Finite_0] addEs [domino.elim]) 1);
 qed "domino_Finite";
 
-Goal "!!d. [| d:domino; b<2 |] ==> EX i' j'. evnodd(d,b) = {<i',j'>}";
+Goal "[| d:domino; b<2 |] ==> EX i' j'. evnodd(d,b) = {<i',j'>}";
 by (eresolve_tac [domino.elim] 1);
 by (res_inst_tac [("k1", "i#+j")] (mod2_cases RS disjE) 2);
 by (res_inst_tac [("k1", "i#+j")] (mod2_cases RS disjE) 1);
 by (REPEAT_FIRST (ares_tac [add_type]));
 (*Four similar cases: case (i#+j) mod 2 = b, 2#-b, ...*)
-by (REPEAT (asm_simp_tac (simpset() addsimps [mod_succ, succ_neq_self] 
-                                   addsplits [split_if]) 1
-           THEN blast_tac (claset() addDs [ltD]) 1));
+by (REPEAT (asm_simp_tac (simpset() addsimps [mod_succ, succ_neq_self]) 1
+	    THEN blast_tac (claset() addDs [ltD]) 1));
 qed "domino_singleton";
 
 
@@ -65,7 +63,7 @@
 
 (** The union of two disjoint tilings is a tiling **)
 
-Goal "!!t. t: tiling(A) ==> \
+Goal "t: tiling(A) ==> \
 \              u: tiling(A) --> t Int u = 0 --> t Un u : tiling(A)";
 by (etac tiling.induct 1);
 by (simp_tac (simpset() addsimps tiling.intrs) 1);
@@ -74,28 +72,30 @@
 by (blast_tac (claset() addIs tiling.intrs) 1);
 qed_spec_mp "tiling_UnI";
 
-Goal "!!t. t:tiling(domino) ==> Finite(t)";
+Goal "t:tiling(domino) ==> Finite(t)";
 by (eresolve_tac [tiling.induct] 1);
 by (rtac Finite_0 1);
 by (blast_tac (claset() addSIs [Finite_Un] addIs [domino_Finite]) 1);
 qed "tiling_domino_Finite";
 
-Goal "!!t. t: tiling(domino) ==> |evnodd(t,0)| = |evnodd(t,1)|";
+Goal "t: tiling(domino) ==> |evnodd(t,0)| = |evnodd(t,1)|";
 by (eresolve_tac [tiling.induct] 1);
 by (simp_tac (simpset() addsimps [evnodd_def]) 1);
 by (res_inst_tac [("b1","0")] (domino_singleton RS exE) 1);
 by (Simp_tac 2 THEN assume_tac 1);
 by (res_inst_tac [("b1","1")] (domino_singleton RS exE) 1);
 by (Simp_tac 2 THEN assume_tac 1);
-by (Step_tac 1);
-by (subgoal_tac "ALL p b. p:evnodd(a,b) --> p~:evnodd(ta,b)" 1);
-by (asm_simp_tac (simpset() addsimps [evnodd_Un, Un_cons, tiling_domino_Finite,
-                                  evnodd_subset RS subset_Finite,
-                                  Finite_imp_cardinal_cons]) 1);
-by (blast_tac (claset() addSDs [evnodd_subset RS subsetD] addEs [equalityE]) 1);
+by Safe_tac;
+by (subgoal_tac "ALL p b. p:evnodd(a,b) --> p~:evnodd(t,b)" 1);
+by (asm_simp_tac 
+    (simpset() addsimps [evnodd_Un, Un_cons, tiling_domino_Finite,
+			 evnodd_subset RS subset_Finite,
+			 Finite_imp_cardinal_cons]) 1);
+by (blast_tac (claset() addSDs [evnodd_subset RS subsetD]
+                        addEs [equalityE]) 1);
 qed "tiling_domino_0_1";
 
-Goal "!!i n. [| i: nat;  n: nat |] ==> {i} * (n #+ n) : tiling(domino)";
+Goal "[| i: nat;  n: nat |] ==> {i} * (n #+ n) : tiling(domino)";
 by (nat_ind_tac "n" [] 1);
 by (simp_tac (simpset() addsimps tiling.intrs) 1);
 by (asm_simp_tac (simpset() addsimps [Un_assoc RS sym, Sigma_succ2]) 1);
@@ -108,7 +108,7 @@
 by (blast_tac (claset() addEs [mem_irrefl, mem_asym]) 1);
 qed "dominoes_tile_row";
 
-Goal "!!m n. [| m: nat;  n: nat |] ==> m * (n #+ n) : tiling(domino)";
+Goal "[| m: nat;  n: nat |] ==> m * (n #+ n) : tiling(domino)";
 by (nat_ind_tac "m" [] 1);
 by (simp_tac (simpset() addsimps tiling.intrs) 1);
 by (asm_simp_tac (simpset() addsimps [Sigma_succ1]) 1);
@@ -117,7 +117,7 @@
 qed "dominoes_tile_matrix";
 
 
-Goal "!!m n. [| m: nat;  n: nat;  \
+Goal "[| m: nat;  n: nat;  \
 \                   t = (succ(m)#+succ(m))*(succ(n)#+succ(n));  \
 \                   t' = t - {<0,0>} - {<succ(m#+m), succ(n#+n)>} |] ==> \
 \                t' ~: tiling(domino)";
--- a/src/ZF/ex/Ntree.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/Ntree.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -48,7 +48,7 @@
 
 (**  Lemmas to justify using "Ntree" in other recursive type definitions **)
 
-Goalw ntree.defs "!!A B. A<=B ==> ntree(A) <= ntree(B)";
+Goalw ntree.defs "A<=B ==> ntree(A) <= ntree(B)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac ntree.bnd_mono 1));
 by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
--- a/src/ZF/ex/Primes.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/Primes.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -14,7 +14,7 @@
 (** Divides Relation                           **)
 (************************************************)
 
-Goalw [dvd_def] "!!m. m dvd n ==> m:nat & n:nat & (EX k:nat. n = m#*k)";
+Goalw [dvd_def] "m dvd n ==> m:nat & n:nat & (EX k:nat. n = m#*k)";
 by (assume_tac 1);
 qed "dvdD";
 
@@ -22,23 +22,23 @@
 bind_thm ("dvd_imp_nat2", dvdD RS conjunct2 RS conjunct1);
 
 
-Goalw [dvd_def] "!!m. m:nat ==> m dvd 0";
+Goalw [dvd_def] "m:nat ==> m dvd 0";
 by (fast_tac (claset() addIs [nat_0I, mult_0_right RS sym]) 1);
 qed "dvd_0_right";
 
-Goalw [dvd_def] "!!m. 0 dvd m ==> m = 0";
+Goalw [dvd_def] "0 dvd m ==> m = 0";
 by (fast_tac (claset() addss (simpset())) 1);
 qed "dvd_0_left";
 
-Goalw [dvd_def] "!!m. m:nat ==> m dvd m";
+Goalw [dvd_def] "m:nat ==> m dvd m";
 by (fast_tac (claset() addIs [nat_1I, mult_1_right RS sym]) 1);
 qed "dvd_refl";
 
-Goalw [dvd_def] "!!m n p. [| m dvd n; n dvd p |] ==> m dvd p";
+Goalw [dvd_def] "[| m dvd n; n dvd p |] ==> m dvd p";
 by (fast_tac (claset() addIs [mult_assoc, mult_type] ) 1);
 qed "dvd_trans";
 
-Goalw [dvd_def] "!!m n. [| m dvd n; n dvd m |] ==> m=n";
+Goalw [dvd_def] "[| m dvd n; n dvd m |] ==> m=n";
 by (fast_tac (claset() addDs [mult_eq_self_implies_10]
                     addss (simpset() addsimps [mult_assoc, mult_eq_1_iff])) 1);
 qed "dvd_anti_sym";
@@ -50,7 +50,7 @@
 
 (* GCD by Euclid's Algorithm *)
 
-Goalw [egcd_def] "!!m. m:nat ==> egcd(m,0) = m";
+Goalw [egcd_def] "m:nat ==> egcd(m,0) = m";
 by (stac transrec 1);
 by (Asm_simp_tac 1);
 qed "egcd_0";
@@ -62,23 +62,23 @@
                                      mod_less_divisor RS ltD]) 1);
 qed "egcd_lt_0";
 
-Goal "!!m. m:nat ==> egcd(m,0) dvd m";
+Goal "m:nat ==> egcd(m,0) dvd m";
 by (asm_simp_tac (simpset() addsimps [egcd_0,dvd_refl]) 1);
 qed "egcd_0_dvd_m";
 
-Goal "!!m. m:nat ==> egcd(m,0) dvd 0";
+Goal "m:nat ==> egcd(m,0) dvd 0";
 by (asm_simp_tac (simpset() addsimps [egcd_0,dvd_0_right]) 1);
 qed "egcd_0_dvd_0";
 
-Goalw [dvd_def] "!!k. [| k dvd a; k dvd b |] ==> k dvd (a #+ b)";
+Goalw [dvd_def] "[| k dvd a; k dvd b |] ==> k dvd (a #+ b)";
 by (fast_tac (claset() addIs [add_mult_distrib_left RS sym, add_type]) 1);
 qed "dvd_add";
 
-Goalw [dvd_def] "!!k. [| k dvd a; q:nat |] ==> k dvd (q #* a)";
+Goalw [dvd_def] "[| k dvd a; q:nat |] ==> k dvd (q #* a)";
 by (fast_tac (claset() addIs [mult_left_commute, mult_type]) 1);
 qed "dvd_mult";
 
-Goal "!!k. [| k dvd b; k dvd (a mod b); 0 < b; a:nat |] ==> k dvd a";
+Goal "[| k dvd b; k dvd (a mod b); 0 < b; a:nat |] ==> k dvd a";
 by (deepen_tac 
     (claset() addIs [mod_div_equality RS subst]
            addDs [dvdD]
@@ -88,7 +88,7 @@
 
 (* egcd type *)
 
-Goal "!!b. b:nat ==> ALL a:nat. egcd(a,b):nat";
+Goal "b:nat ==> ALL a:nat. egcd(a,b):nat";
 by (etac complete_induct 1);
 by (rtac ballI 1);
 by (excluded_middle_tac "x=0" 1);
@@ -105,7 +105,7 @@
 
 (* Property 1: egcd(a,b) divides a and b *)
 
-Goal "!!b. b:nat ==> ALL a: nat. (egcd(a,b) dvd a) & (egcd(a,b) dvd b)";
+Goal "b:nat ==> ALL a: nat. (egcd(a,b) dvd a) & (egcd(a,b) dvd b)";
 by (res_inst_tac [("i","b")] complete_induct 1);
 by (assume_tac 1);
 by (rtac ballI 1);
@@ -124,7 +124,7 @@
 
 (* if f divides a and b then f divides egcd(a,b) *)
 
-Goalw [dvd_def] "!!a. [| f dvd a; f dvd b; 0<b |] ==> f dvd (a mod b)";
+Goalw [dvd_def] "[| f dvd a; f dvd b; 0<b |] ==> f dvd (a mod b)";
 by (safe_tac (claset() addSIs [mult_type, mod_type]));
 ren "m n" 1;
 by (rtac (zero_lt_mult_iff RS iffD1 RS conjE) 1);
@@ -141,7 +141,7 @@
 
 (* Property 2: for all a,b,f naturals, 
                if f divides a and f divides b then f divides egcd(a,b)*)
-Goal "!!b. [| b:nat; f:nat |] ==>    \
+Goal "[| b:nat; f:nat |] ==>    \
 \              ALL a. (f dvd a) & (f dvd b) --> f dvd egcd(a,b)";
 by (etac complete_induct 1);
 by (rtac allI 1);
@@ -162,14 +162,14 @@
 
 (* GCD PROOF : GCD exists and egcd fits the definition *)
 
-Goalw [gcd_def] "!!b. [| a: nat; b:nat |] ==> gcd(egcd(a,b), a, b)";
+Goalw [gcd_def] "[| a: nat; b:nat |] ==> gcd(egcd(a,b), a, b)";
 by (asm_simp_tac (simpset() addsimps [egcd_prop1]) 1);
 by (fast_tac (claset() addIs [egcd_prop2 RS spec RS mp, dvd_imp_nat1]) 1);
 qed "gcd";
 
 (* GCD is unique *)
 
-Goalw [gcd_def] "!!a. gcd(m,a,b) & gcd(n,a,b) --> m=n";
+Goalw [gcd_def] "gcd(m,a,b) & gcd(n,a,b) --> m=n";
 by (fast_tac (claset() addIs [dvd_anti_sym]) 1);
 qed "gcd_unique";
 
--- a/src/ZF/ex/Primrec.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/Primrec.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -64,7 +64,7 @@
 simpset_ref() := simpset() setSolver (type_auto_tac ([primrec_into_fun] @ 
 					      pr_typechecks @ primrec.intrs));
 
-Goalw [ACK_def] "!!i. i:nat ==> ACK(i): primrec";
+Goalw [ACK_def] "i:nat ==> ACK(i): primrec";
 by (etac nat_induct 1);
 by (ALLGOALS Asm_simp_tac);
 qed "ACK_in_primrec";
@@ -79,14 +79,14 @@
     REPEAT
       (SOMEGOAL (test_assume_tac ORELSE' match_tac (rls @ ack_typechecks)));
 
-Goal "!!i j. [| i:nat;  j:nat |] ==>  ack(i,j): nat";
+Goal "[| i:nat;  j:nat |] ==>  ack(i,j): nat";
 by (tc_tac []);
 qed "ack_type";
 
 (** Ackermann's function cases **)
 
 (*PROPERTY A 1*)
-Goalw [ACK_def] "!!j. j:nat ==> ack(0,j) = succ(j)";
+Goalw [ACK_def] "j:nat ==> ack(0,j) = succ(j)";
 by (asm_simp_tac (simpset() addsimps [SC]) 1);
 qed "ack_0";
 
@@ -107,7 +107,7 @@
 Addsimps [ack_0, ack_succ_0, ack_succ_succ, ack_type, nat_into_Ord];
 
 (*PROPERTY A 4*)
-Goal "!!i. i:nat ==> ALL j:nat. j < ack(i,j)";
+Goal "i:nat ==> ALL j:nat. j < ack(i,j)";
 by (etac nat_induct 1);
 by (Asm_simp_tac 1);
 by (rtac ballI 1);
@@ -120,13 +120,13 @@
 bind_thm ("lt_ack2", (lt_ack2_lemma RS bspec));
 
 (*PROPERTY A 5-, the single-step lemma*)
-Goal "!!i j. [| i:nat; j:nat |] ==> ack(i,j) < ack(i, succ(j))";
+Goal "[| i:nat; j:nat |] ==> ack(i,j) < ack(i, succ(j))";
 by (etac nat_induct 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [lt_ack2])));
 qed "ack_lt_ack_succ2";
 
 (*PROPERTY A 5, monotonicity for < *)
-Goal "!!i j k. [| j<k; i:nat; k:nat |] ==> ack(i,j) < ack(i,k)";
+Goal "[| j<k; i:nat; k:nat |] ==> ack(i,j) < ack(i,k)";
 by (forward_tac [lt_nat_in_nat] 1 THEN assume_tac 1);
 by (etac succ_lt_induct 1);
 by (assume_tac 1);
@@ -152,14 +152,14 @@
 qed "ack2_le_ack1";
 
 (*PROPERTY A 7-, the single-step lemma*)
-Goal "!!i j. [| i:nat; j:nat |] ==> ack(i,j) < ack(succ(i),j)";
+Goal "[| i:nat; j:nat |] ==> ack(i,j) < ack(succ(i),j)";
 by (rtac (ack_lt_mono2 RS lt_trans2) 1);
 by (rtac ack2_le_ack1 4);
 by (REPEAT (ares_tac ([nat_le_refl, ack_type] @ pr_typechecks) 1));
 qed "ack_lt_ack_succ1";
 
 (*PROPERTY A 7, monotonicity for < *)
-Goal "!!i j k. [| i<j; j:nat; k:nat |] ==> ack(i,k) < ack(j,k)";
+Goal "[| i<j; j:nat; k:nat |] ==> ack(i,k) < ack(j,k)";
 by (forward_tac [lt_nat_in_nat] 1 THEN assume_tac 1);
 by (etac succ_lt_induct 1);
 by (assume_tac 1);
@@ -175,13 +175,13 @@
 qed "ack_le_mono1";
 
 (*PROPERTY A 8*)
-Goal "!!j. j:nat ==> ack(1,j) = succ(succ(j))";
+Goal "j:nat ==> ack(1,j) = succ(succ(j))";
 by (etac nat_induct 1);
 by (ALLGOALS Asm_simp_tac);
 qed "ack_1";
 
 (*PROPERTY A 9*)
-Goal "!!j. j:nat ==> ack(succ(1),j) = succ(succ(succ(j#+j)))";
+Goal "j:nat ==> ack(succ(1),j) = succ(succ(succ(j#+j)))";
 by (etac nat_induct 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [ack_1, add_succ_right])));
 qed "ack_2";
@@ -233,7 +233,7 @@
 qed "SC_case";
 
 (*PROPERTY A 4'? Extra lemma needed for CONST case, constant functions*)
-Goal "!!j. [| i:nat; j:nat |] ==> i < ack(i,j)";
+Goal "[| i:nat; j:nat |] ==> i < ack(i,j)";
 by (etac nat_induct 1);
 by (asm_simp_tac (simpset() addsimps [nat_0_le]) 1);
 by (etac ([succ_leI, ack_lt_ack_succ1] MRS lt_trans1) 1);
--- a/src/ZF/ex/PropLog.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/PropLog.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -74,7 +74,7 @@
 
 (*** Proof theory of propositional logic ***)
 
-Goalw thms.defs "!!G H. G<=H ==> thms(G) <= thms(H)";
+Goalw thms.defs "G<=H ==> thms(G) <= thms(H)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac thms.bnd_mono 1));
 by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
@@ -85,13 +85,13 @@
 val ImpE = prop.mk_cases prop.con_defs "p=>q : prop";
 
 (*Stronger Modus Ponens rule: no typechecking!*)
-Goal "!!p q H. [| H |- p=>q;  H |- p |] ==> H |- q";
+Goal "[| H |- p=>q;  H |- p |] ==> H |- q";
 by (rtac thms.MP 1);
 by (REPEAT (eresolve_tac [asm_rl, thms_in_pl, thms_in_pl RS ImpE] 1));
 qed "thms_MP";
 
 (*Rule is called I for Identity Combinator, not for Introduction*)
-Goal "!!p H. p:prop ==> H |- p=>p";
+Goal "p:prop ==> H |- p=>p";
 by (rtac (thms.S RS thms_MP RS thms_MP) 1);
 by (rtac thms.K 5);
 by (rtac thms.K 4);
@@ -109,13 +109,13 @@
 val weaken_left_Un1  = Un_upper1 RS weaken_left;
 val weaken_left_Un2  = Un_upper2 RS weaken_left;
 
-Goal "!!H p q. [| H |- q;  p:prop |] ==> H |- p=>q";
+Goal "[| H |- q;  p:prop |] ==> H |- p=>q";
 by (rtac (thms.K RS thms_MP) 1);
 by (REPEAT (ares_tac [thms_in_pl] 1));
 qed "weaken_right";
 
 (*The deduction theorem*)
-Goal "!!p q H. [| cons(p,H) |- q;  p:prop |] ==>  H |- p=>q";
+Goal "[| cons(p,H) |- q;  p:prop |] ==>  H |- p=>q";
 by (etac thms.induct 1);
 by (fast_tac (claset() addIs [thms_I, thms.H RS weaken_right]) 1);
 by (fast_tac (claset() addIs [thms.K RS weaken_right]) 1);
@@ -126,12 +126,12 @@
 
 
 (*The cut rule*)
-Goal "!!H p q. [| H|-p;  cons(p,H) |- q |] ==>  H |- q";
+Goal "[| H|-p;  cons(p,H) |- q |] ==>  H |- q";
 by (rtac (deduction RS thms_MP) 1);
 by (REPEAT (ares_tac [thms_in_pl] 1));
 qed "cut";
 
-Goal "!!H p. [| H |- Fls; p:prop |] ==> H |- p";
+Goal "[| H |- Fls; p:prop |] ==> H |- p";
 by (rtac (thms.DN RS thms_MP) 1);
 by (rtac weaken_right 2);
 by (REPEAT (ares_tac (prop.intrs@[consI1]) 1));
@@ -141,7 +141,7 @@
 bind_thm ("thms_notE", (thms_MP RS thms_FlsE));
 
 (*Soundness of the rules wrt truth-table semantics*)
-Goalw [logcon_def] "!!H. H |- p ==> H |= p";
+Goalw [logcon_def] "H |- p ==> H |= p";
 by (etac thms.induct 1);
 by (fast_tac (claset() addSDs [is_true_Imp RS iffD1 RS mp]) 5);
 by (ALLGOALS Asm_simp_tac);
@@ -289,17 +289,17 @@
 qed "completeness_0";
 
 (*A semantic analogue of the Deduction Theorem*)
-Goalw [logcon_def] "!!H p q. [| cons(p,H) |= q |] ==> H |= p=>q";
+Goalw [logcon_def] "[| cons(p,H) |= q |] ==> H |= p=>q";
 by (Simp_tac 1);
 by (Fast_tac 1);
 qed "logcon_Imp";
 
-Goal "!!H. H: Fin(prop) ==> ALL p:prop. H |= p --> H |- p";
+Goal "H: Fin(prop) ==> ALL p:prop. H |= p --> H |- p";
 by (etac Fin_induct 1);
 by (safe_tac (claset() addSIs [completeness_0]));
 by (rtac (weaken_left_cons RS thms_MP) 1);
-by (fast_tac (claset() addSIs (logcon_Imp::prop.intrs)) 1);
-by (fast_tac thms_cs 1);
+by (blast_tac (claset() addSIs (logcon_Imp::prop.intrs)) 1);
+by (blast_tac thms_cs 1);
 qed "completeness_lemma";
 
 val completeness = completeness_lemma RS bspec RS mp;
--- a/src/ZF/ex/Ramsey.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/Ramsey.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -197,12 +197,12 @@
 by (nat_ind_tac "j" [] 1);
 by (fast_tac (claset() addSIs [Ramseyi0]) 1);
 by (fast_tac (claset() addSDs [bspec]
-		      addSIs [nat_succI,add_type,Ramsey_step_lemma]) 1);
+ 		       addSIs [add_type,Ramsey_step_lemma]) 1);
 qed "ramsey_lemma";
 
 (*Final statement in a tidy form, without succ(...) *)
-Goal "!!i j. [| i: nat;  j: nat |] ==> EX n:nat. Ramsey(n,i,j)";
-by (best_tac (claset() addDs [ramsey_lemma] addSIs [nat_succI]) 1);
+Goal "[| i: nat;  j: nat |] ==> EX n:nat. Ramsey(n,i,j)";
+by (best_tac (claset() addDs [ramsey_lemma]) 1);
 qed "ramsey";
 
 (*Compute Ramsey numbers according to proof above -- which, actually,
--- a/src/ZF/ex/Rmap.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/Rmap.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -8,7 +8,7 @@
 
 open Rmap;
 
-Goalw rmap.defs "!!r s. r<=s ==> rmap(r) <= rmap(s)";
+Goalw rmap.defs "r<=s ==> rmap(r) <= rmap(s)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac rmap.bnd_mono 1));
 by (REPEAT (ares_tac ([Sigma_mono, list_mono, domain_mono, range_mono] @ 
@@ -21,7 +21,7 @@
 AddIs  rmap.intrs;
 AddSEs [Nil_rmap_case, Cons_rmap_case];
 
-Goal "!!r. r <= A*B ==> rmap(r) <= list(A)*list(B)";
+Goal "r <= A*B ==> rmap(r) <= list(A)*list(B)";
 by (rtac (rmap.dom_subset RS subset_trans) 1);
 by (REPEAT (ares_tac [domain_rel_subset, range_rel_subset,
                       Sigma_mono, list_mono] 1));
@@ -34,7 +34,7 @@
 by (ALLGOALS Fast_tac);
 qed "rmap_total";
 
-Goalw [function_def] "!!r. function(r) ==> function(rmap(r))";
+Goalw [function_def] "function(r) ==> function(rmap(r))";
 by (rtac (impI RS allI RS allI) 1);
 by (etac rmap.induct 1);
 by (ALLGOALS Fast_tac);
@@ -43,7 +43,7 @@
 
 (** If f is a function then rmap(f) behaves as expected. **)
 
-Goal "!!f. f: A->B ==> rmap(f): list(A)->list(B)";
+Goal "f: A->B ==> rmap(f): list(A)->list(B)";
 by (asm_full_simp_tac 
     (simpset() addsimps [Pi_iff, rmap_rel_type, rmap_functional, rmap_total]) 1);
 qed "rmap_fun_type";
@@ -52,7 +52,7 @@
 by (fast_tac (claset() addIs [the_equality]) 1);
 qed "rmap_Nil";
 
-Goal "!!f. [| f: A->B;  x: A;  xs: list(A) |] ==> \
+Goal "[| f: A->B;  x: A;  xs: list(A) |] ==> \
 \                   rmap(f) ` Cons(x,xs) = Cons(f`x, rmap(f)`xs)";
 by (rtac apply_equality 1);
 by (REPEAT (ares_tac ([apply_Pair, rmap_fun_type] @ rmap.intrs) 1));
--- a/src/ZF/ex/TF.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/TF.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -216,7 +216,7 @@
 by (REPEAT (ares_tac (TrueI::prems) 1));
 qed "forest_induct";
 
-Goal "!!f A. f: forest(A) ==> TF_of_list(list_of_TF(f)) = f";
+Goal "f: forest(A) ==> TF_of_list(list_of_TF(f)) = f";
 by (etac forest_induct 1);
 by (ALLGOALS Asm_simp_tac);
 qed "forest_iso";
@@ -229,7 +229,7 @@
 
 (** theorems about TF_map **)
 
-Goal "!!z A. z: tree_forest(A) ==> TF_map(%u. u, z) = z";
+Goal "z: tree_forest(A) ==> TF_map(%u. u, z) = z";
 by (etac tree_forest.induct 1);
 by (ALLGOALS Asm_simp_tac);
 qed "TF_map_ident";
@@ -256,7 +256,7 @@
 
 (** theorems about TF_preorder **)
 
-Goal "!!z A. z: tree_forest(A) ==> \
+Goal "z: tree_forest(A) ==> \
 \                      TF_preorder(TF_map(h,z)) = map(h, TF_preorder(z))";
 by (etac tree_forest.induct 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [map_app_distrib])));
--- a/src/ZF/ex/Term.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/ex/Term.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -41,7 +41,7 @@
 
 (**  Lemmas to justify using "term" in other recursive type definitions **)
 
-Goalw term.defs "!!A B. A<=B ==> term(A) <= term(B)";
+Goalw term.defs "A<=B ==> term(A) <= term(B)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac term.bnd_mono 1));
 by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
@@ -58,7 +58,7 @@
 val term_subset_univ = 
     term_mono RS (term_univ RSN (2,subset_trans)) |> standard;
 
-Goal "!!t A B. [| t: term(A);  A <= univ(B) |] ==> t: univ(B)";
+Goal "[| t: term(A);  A <= univ(B) |] ==> t: univ(B)";
 by (REPEAT (ares_tac [term_subset_univ RS subsetD] 1));
 qed "term_into_univ";
 
@@ -140,7 +140,7 @@
 
 bind_thm ("term_size", (term_size_def RS def_term_rec));
 
-Goalw [term_size_def] "!!t A. t: term(A) ==> term_size(t) : nat";
+Goalw [term_size_def] "t: term(A) ==> term_size(t) : nat";
 by (REPEAT (ares_tac [term_rec_simple_type, list_add_type, nat_succI] 1));
 qed "term_size_type";
 
@@ -149,7 +149,7 @@
 
 bind_thm ("reflect", (reflect_def RS def_term_rec));
 
-Goalw [reflect_def] "!!t A. t: term(A) ==> reflect(t) : term(A)";
+Goalw [reflect_def] "t: term(A) ==> reflect(t) : term(A)";
 by (REPEAT (ares_tac [term_rec_simple_type, rev_type, term.Apply_I] 1));
 qed "reflect_type";
 
@@ -177,7 +177,7 @@
 
 (** theorems about term_map **)
 
-Goal "!!t A. t: term(A) ==> term_map(%u. u, t) = t";
+Goal "t: term(A) ==> term_map(%u. u, t) = t";
 by (etac term_induct_eqn 1);
 by (asm_simp_tac (simpset() addsimps [map_ident]) 1);
 qed "term_map_ident";
@@ -203,13 +203,13 @@
 by (asm_simp_tac (simpset() addsimps [map_compose]) 1);
 qed "term_size_term_map";
 
-Goal "!!t A. t: term(A) ==> term_size(reflect(t)) = term_size(t)";
+Goal "t: term(A) ==> term_size(reflect(t)) = term_size(t)";
 by (etac term_induct_eqn 1);
 by (asm_simp_tac (simpset() addsimps [rev_map_distrib RS sym, map_compose,
                                     list_add_rev]) 1);
 qed "term_size_reflect";
 
-Goal "!!t A. t: term(A) ==> term_size(t) = length(preorder(t))";
+Goal "t: term(A) ==> term_size(t) = length(preorder(t))";
 by (etac term_induct_eqn 1);
 by (asm_simp_tac (simpset() addsimps [length_flat, map_compose]) 1);
 qed "term_size_length";
@@ -217,7 +217,7 @@
 
 (** theorems about reflect **)
 
-Goal "!!t A. t: term(A) ==> reflect(reflect(t)) = t";
+Goal "t: term(A) ==> reflect(reflect(t)) = t";
 by (etac term_induct_eqn 1);
 by (asm_simp_tac (simpset() addsimps [rev_map_distrib, map_compose,
                                     map_ident, rev_rev_ident]) 1);
--- a/src/ZF/func.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/func.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -253,7 +253,7 @@
                               addcongs [RepFun_cong]) 1);
 qed "image_fun";
 
-Goal "!!f. [| f: Pi(A,B);  x: A |] ==> f `` cons(x,y) = cons(f`x, f``y)";
+Goal "[| f: Pi(A,B);  x: A |] ==> f `` cons(x,y) = cons(f`x, f``y)";
 by (blast_tac (claset() addDs [apply_equality, apply_Pair]) 1);
 qed "Pi_image_cons";
 
--- a/src/ZF/mono.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/mono.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -12,19 +12,19 @@
 
 (*Not easy to express monotonicity in P, since any "bigger" predicate
   would have to be single-valued*)
-Goal "!!A B. A<=B ==> Replace(A,P) <= Replace(B,P)";
+Goal "A<=B ==> Replace(A,P) <= Replace(B,P)";
 by (blast_tac (claset() addSEs [ReplaceE]) 1);
 qed "Replace_mono";
 
-Goal "!!A B. A<=B ==> {f(x). x:A} <= {f(x). x:B}";
+Goal "A<=B ==> {f(x). x:A} <= {f(x). x:B}";
 by (Blast_tac 1);
 qed "RepFun_mono";
 
-Goal "!!A B. A<=B ==> Pow(A) <= Pow(B)";
+Goal "A<=B ==> Pow(A) <= Pow(B)";
 by (Blast_tac 1);
 qed "Pow_mono";
 
-Goal "!!A B. A<=B ==> Union(A) <= Union(B)";
+Goal "A<=B ==> Union(A) <= Union(B)";
 by (Blast_tac 1);
 qed "Union_mono";
 
@@ -35,95 +35,95 @@
 qed "UN_mono";
 
 (*Intersection is ANTI-monotonic.  There are TWO premises! *)
-Goal "!!A B. [| A<=B;  a:A |] ==> Inter(B) <= Inter(A)";
+Goal "[| A<=B;  a:A |] ==> Inter(B) <= Inter(A)";
 by (Blast_tac 1);
 qed "Inter_anti_mono";
 
-Goal "!!C D. C<=D ==> cons(a,C) <= cons(a,D)";
+Goal "C<=D ==> cons(a,C) <= cons(a,D)";
 by (Blast_tac 1);
 qed "cons_mono";
 
-Goal "!!A B C D. [| A<=C;  B<=D |] ==> A Un B <= C Un D";
+Goal "[| A<=C;  B<=D |] ==> A Un B <= C Un D";
 by (Blast_tac 1);
 qed "Un_mono";
 
-Goal "!!A B C D. [| A<=C;  B<=D |] ==> A Int B <= C Int D";
+Goal "[| A<=C;  B<=D |] ==> A Int B <= C Int D";
 by (Blast_tac 1);
 qed "Int_mono";
 
-Goal "!!A B C D. [| A<=C;  D<=B |] ==> A-B <= C-D";
+Goal "[| A<=C;  D<=B |] ==> A-B <= C-D";
 by (Blast_tac 1);
 qed "Diff_mono";
 
 (** Standard products, sums and function spaces **)
 
-Goal "!!A B C D. [| A<=C;  ALL x:A. B(x) <= D(x) |] ==> \
+Goal "[| A<=C;  ALL x:A. B(x) <= D(x) |] ==> \
 \                       Sigma(A,B) <= Sigma(C,D)";
 by (Blast_tac 1);
 qed "Sigma_mono_lemma";
 val Sigma_mono = ballI RSN (2,Sigma_mono_lemma);
 
-Goalw sum_defs "!!A B C D. [| A<=C;  B<=D |] ==> A+B <= C+D";
+Goalw sum_defs "[| A<=C;  B<=D |] ==> A+B <= C+D";
 by (REPEAT (ares_tac [subset_refl,Un_mono,Sigma_mono] 1));
 qed "sum_mono";
 
 (*Note that B->A and C->A are typically disjoint!*)
-Goal "!!A B C. B<=C ==> A->B <= A->C";
+Goal "B<=C ==> A->B <= A->C";
 by (blast_tac (claset() addIs [lam_type] addEs [Pi_lamE]) 1);
 qed "Pi_mono";
 
-Goalw [lam_def] "!!A B. A<=B ==> Lambda(A,c) <= Lambda(B,c)";
+Goalw [lam_def] "A<=B ==> Lambda(A,c) <= Lambda(B,c)";
 by (etac RepFun_mono 1);
 qed "lam_mono";
 
 (** Quine-inspired ordered pairs, products, injections and sums **)
 
-Goalw [QPair_def] "!!a b c d. [| a<=c;  b<=d |] ==> <a;b> <= <c;d>";
+Goalw [QPair_def] "[| a<=c;  b<=d |] ==> <a;b> <= <c;d>";
 by (REPEAT (ares_tac [sum_mono] 1));
 qed "QPair_mono";
 
-Goal "!!A B C D. [| A<=C;  ALL x:A. B(x) <= D(x) |] ==>  \
+Goal "[| A<=C;  ALL x:A. B(x) <= D(x) |] ==>  \
 \                          QSigma(A,B) <= QSigma(C,D)";
 by (Blast_tac 1);
 qed "QSigma_mono_lemma";
 val QSigma_mono = ballI RSN (2,QSigma_mono_lemma);
 
-Goalw [QInl_def] "!!a b. a<=b ==> QInl(a) <= QInl(b)";
+Goalw [QInl_def] "a<=b ==> QInl(a) <= QInl(b)";
 by (REPEAT (ares_tac [subset_refl RS QPair_mono] 1));
 qed "QInl_mono";
 
-Goalw [QInr_def] "!!a b. a<=b ==> QInr(a) <= QInr(b)";
+Goalw [QInr_def] "a<=b ==> QInr(a) <= QInr(b)";
 by (REPEAT (ares_tac [subset_refl RS QPair_mono] 1));
 qed "QInr_mono";
 
-Goal "!!A B C D. [| A<=C;  B<=D |] ==> A <+> B <= C <+> D";
+Goal "[| A<=C;  B<=D |] ==> A <+> B <= C <+> D";
 by (Blast_tac 1);
 qed "qsum_mono";
 
 
 (** Converse, domain, range, field **)
 
-Goal "!!r s. r<=s ==> converse(r) <= converse(s)";
+Goal "r<=s ==> converse(r) <= converse(s)";
 by (Blast_tac 1);
 qed "converse_mono";
 
-Goal "!!r s. r<=s ==> domain(r)<=domain(s)";
+Goal "r<=s ==> domain(r)<=domain(s)";
 by (Blast_tac 1);
 qed "domain_mono";
 
 bind_thm ("domain_rel_subset", [domain_mono, domain_subset] MRS subset_trans);
 
-Goal "!!r s. r<=s ==> range(r)<=range(s)";
+Goal "r<=s ==> range(r)<=range(s)";
 by (Blast_tac 1);
 qed "range_mono";
 
 bind_thm ("range_rel_subset", [range_mono, range_subset] MRS subset_trans);
 
-Goal "!!r s. r<=s ==> field(r)<=field(s)";
+Goal "r<=s ==> field(r)<=field(s)";
 by (Blast_tac 1);
 qed "field_mono";
 
-Goal "!!r A. r <= A*A ==> field(r) <= A";
+Goal "r <= A*A ==> field(r) <= A";
 by (etac (field_mono RS subset_trans) 1);
 by (Blast_tac 1);
 qed "field_rel_subset";
@@ -141,11 +141,11 @@
 by (blast_tac (claset() addIs [prem1, prem2 RS subsetD]) 1);
 qed "vimage_pair_mono";
 
-Goal "!!r s. [| r<=s;  A<=B |] ==> r``A <= s``B";
+Goal "[| r<=s;  A<=B |] ==> r``A <= s``B";
 by (Blast_tac 1);
 qed "image_mono";
 
-Goal "!!r s. [| r<=s;  A<=B |] ==> r-``A <= s-``B";
+Goal "[| r<=s;  A<=B |] ==> r-``A <= s-``B";
 by (Blast_tac 1);
 qed "vimage_mono";
 
@@ -156,7 +156,7 @@
 
 (** Monotonicity of implications -- some could go to FOL **)
 
-Goal "!!A B x. A<=B ==> x:A --> x:B";
+Goal "A<=B ==> x:A --> x:B";
 by (Blast_tac 1);
 qed "in_mono";
 
--- a/src/ZF/simpdata.ML	Mon Jul 13 16:42:27 1998 +0200
+++ b/src/ZF/simpdata.ML	Mon Jul 13 16:43:57 1998 +0200
@@ -104,6 +104,7 @@
 
 val ZF_atomize = atomize (ZF_conn_pairs, ZF_mem_pairs);
 
-simpset_ref() := simpset() setmksimps (map mk_meta_eq o ZF_atomize o gen_all);
+simpset_ref() := simpset() setmksimps (map mk_meta_eq o ZF_atomize o gen_all)
+                           addsplits [split_if];
 
 val ZF_ss = simpset();