--- a/src/ZF/AC.ML Mon Jan 29 14:16:13 1996 +0100
+++ b/src/ZF/AC.ML Tue Jan 30 13:42:57 1996 +0100
@@ -1,6 +1,6 @@
-(* Title: ZF/AC.ML
+(* Title: ZF/AC.ML
ID: $Id$
- Author: Lawrence C Paulson, Cambridge University Computer Laboratory
+ Author: Lawrence C Paulson, Cambridge University Computer Laboratory
Copyright 1994 University of Cambridge
For AC.thy. The Axiom of Choice
@@ -19,8 +19,8 @@
(*Using dtac, this has the advantage of DELETING the universal quantifier*)
goal AC.thy "!!A B. ALL x:A. EX y. y:B(x) ==> EX y. y : Pi(A,B)";
-by (resolve_tac [AC_Pi] 1);
-by (eresolve_tac [bspec] 1);
+by (rtac AC_Pi 1);
+by (etac bspec 1);
by (assume_tac 1);
qed "AC_ball_Pi";
@@ -31,7 +31,7 @@
qed "AC_Pi_Pow";
val [nonempty] = goal AC.thy
- "[| !!x. x:A ==> (EX y. y:x) \
+ "[| !!x. x:A ==> (EX y. y:x) \
\ |] ==> EX f: A->Union(A). ALL x:A. f`x : x";
by (res_inst_tac [("B1", "%x.x")] (AC_Pi RS exE) 1);
by (etac nonempty 1);
@@ -52,7 +52,7 @@
by (resolve_tac [AC_func0 RS bexE] 1);
by (rtac bexI 2);
by (assume_tac 2);
-by (eresolve_tac [fun_weaken_type] 2);
+by (etac fun_weaken_type 2);
by (ALLGOALS (fast_tac ZF_cs));
qed "AC_func_Pow";