| author | wenzelm | 
| Fri, 10 Oct 1997 19:13:58 +0200 | |
| changeset 3843 | 162f95673705 | 
| parent 3840 | e0baea4d485a | 
| child 4091 | 771b1f6422a8 | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/AC.ML | 
| 484 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 484 | 4 | Copyright 1994 University of Cambridge | 
| 5 | ||
| 6 | For AC.thy. The Axiom of Choice | |
| 7 | *) | |
| 8 | ||
| 9 | open AC; | |
| 10 | ||
| 11 | (*The same as AC, but no premise a:A*) | |
| 12 | val [nonempty] = goal AC.thy | |
| 13 | "[| !!x. x:A ==> (EX y. y:B(x)) |] ==> EX z. z : Pi(A,B)"; | |
| 14 | by (excluded_middle_tac "A=0" 1); | |
| 3016 | 15 | by (asm_simp_tac (!simpset addsimps [Pi_empty1]) 2 THEN Blast_tac 2); | 
| 484 | 16 | (*The non-trivial case*) | 
| 3016 | 17 | by (blast_tac (!claset addIs [AC, nonempty]) 1); | 
| 760 | 18 | qed "AC_Pi"; | 
| 484 | 19 | |
| 20 | (*Using dtac, this has the advantage of DELETING the universal quantifier*) | |
| 21 | goal AC.thy "!!A B. ALL x:A. EX y. y:B(x) ==> EX y. y : Pi(A,B)"; | |
| 1461 | 22 | by (rtac AC_Pi 1); | 
| 23 | by (etac bspec 1); | |
| 484 | 24 | by (assume_tac 1); | 
| 760 | 25 | qed "AC_ball_Pi"; | 
| 484 | 26 | |
| 27 | goal AC.thy "EX f. f: (PROD X: Pow(C)-{0}. X)";
 | |
| 3840 | 28 | by (res_inst_tac [("B1", "%x. x")] (AC_Pi RS exE) 1);
 | 
| 484 | 29 | by (etac exI 2); | 
| 3016 | 30 | by (Blast_tac 1); | 
| 760 | 31 | qed "AC_Pi_Pow"; | 
| 484 | 32 | |
| 33 | val [nonempty] = goal AC.thy | |
| 1461 | 34 | "[| !!x. x:A ==> (EX y. y:x) \ | 
| 484 | 35 | \ |] ==> EX f: A->Union(A). ALL x:A. f`x : x"; | 
| 3840 | 36 | by (res_inst_tac [("B1", "%x. x")] (AC_Pi RS exE) 1);
 | 
| 484 | 37 | by (etac nonempty 1); | 
| 3016 | 38 | by (blast_tac (!claset addDs [apply_type] addIs [Pi_type]) 1); | 
| 760 | 39 | qed "AC_func"; | 
| 484 | 40 | |
| 741 | 41 | goal ZF.thy "!!x A. [| 0 ~: A; x: A |] ==> EX y. y:x"; | 
| 42 | by (subgoal_tac "x ~= 0" 1); | |
| 3016 | 43 | by (ALLGOALS Blast_tac); | 
| 760 | 44 | qed "non_empty_family"; | 
| 484 | 45 | |
| 46 | goal AC.thy "!!A. 0 ~: A ==> EX f: A->Union(A). ALL x:A. f`x : x"; | |
| 47 | by (rtac AC_func 1); | |
| 48 | by (REPEAT (ares_tac [non_empty_family] 1)); | |
| 760 | 49 | qed "AC_func0"; | 
| 484 | 50 | |
| 51 | goal AC.thy "EX f: (Pow(C)-{0}) -> C. ALL x:(Pow(C)-{0}). f`x : x";
 | |
| 52 | by (resolve_tac [AC_func0 RS bexE] 1); | |
| 53 | by (rtac bexI 2); | |
| 54 | by (assume_tac 2); | |
| 1461 | 55 | by (etac fun_weaken_type 2); | 
| 3016 | 56 | by (ALLGOALS Blast_tac); | 
| 760 | 57 | qed "AC_func_Pow"; | 
| 484 | 58 | |
| 1074 
d60f203eeddf
Modified proofs for new claset primitives.  The problem is that they enforce
 lcp parents: 
760diff
changeset | 59 | goal AC.thy "!!A. 0 ~: A ==> EX f. f: (PROD x:A. x)"; | 
| 
d60f203eeddf
Modified proofs for new claset primitives.  The problem is that they enforce
 lcp parents: 
760diff
changeset | 60 | by (rtac AC_Pi 1); | 
| 
d60f203eeddf
Modified proofs for new claset primitives.  The problem is that they enforce
 lcp parents: 
760diff
changeset | 61 | by (REPEAT (ares_tac [non_empty_family] 1)); | 
| 
d60f203eeddf
Modified proofs for new claset primitives.  The problem is that they enforce
 lcp parents: 
760diff
changeset | 62 | qed "AC_Pi0"; | 
| 
d60f203eeddf
Modified proofs for new claset primitives.  The problem is that they enforce
 lcp parents: 
760diff
changeset | 63 |