| author | wenzelm | 
| Fri, 10 Nov 2000 19:02:37 +0100 | |
| changeset 10432 | 3dfbc913d184 | 
| parent 9907 | 473a6604da94 | 
| child 11320 | 56aa53caf333 | 
| 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 | ||
| 5809 | 6 | The Axiom of Choice | 
| 484 | 7 | *) | 
| 8 | ||
| 9 | (*The same as AC, but no premise a:A*) | |
| 9907 | 10 | val [nonempty] = goal (the_context ()) | 
| 484 | 11 | "[| !!x. x:A ==> (EX y. y:B(x)) |] ==> EX z. z : Pi(A,B)"; | 
| 12 | by (excluded_middle_tac "A=0" 1); | |
| 4091 | 13 | by (asm_simp_tac (simpset() addsimps [Pi_empty1]) 2 THEN Blast_tac 2); | 
| 484 | 14 | (*The non-trivial case*) | 
| 4091 | 15 | by (blast_tac (claset() addIs [AC, nonempty]) 1); | 
| 760 | 16 | qed "AC_Pi"; | 
| 484 | 17 | |
| 18 | (*Using dtac, this has the advantage of DELETING the universal quantifier*) | |
| 5137 | 19 | Goal "ALL x:A. EX y. y:B(x) ==> EX y. y : Pi(A,B)"; | 
| 1461 | 20 | by (rtac AC_Pi 1); | 
| 21 | by (etac bspec 1); | |
| 484 | 22 | by (assume_tac 1); | 
| 760 | 23 | qed "AC_ball_Pi"; | 
| 484 | 24 | |
| 5067 | 25 | Goal "EX f. f: (PROD X: Pow(C)-{0}. X)";
 | 
| 3840 | 26 | by (res_inst_tac [("B1", "%x. x")] (AC_Pi RS exE) 1);
 | 
| 484 | 27 | by (etac exI 2); | 
| 3016 | 28 | by (Blast_tac 1); | 
| 760 | 29 | qed "AC_Pi_Pow"; | 
| 484 | 30 | |
| 9907 | 31 | val [nonempty] = goal (the_context ()) | 
| 1461 | 32 | "[| !!x. x:A ==> (EX y. y:x) \ | 
| 484 | 33 | \ |] ==> EX f: A->Union(A). ALL x:A. f`x : x"; | 
| 3840 | 34 | by (res_inst_tac [("B1", "%x. x")] (AC_Pi RS exE) 1);
 | 
| 484 | 35 | by (etac nonempty 1); | 
| 4091 | 36 | by (blast_tac (claset() addDs [apply_type] addIs [Pi_type]) 1); | 
| 760 | 37 | qed "AC_func"; | 
| 484 | 38 | |
| 741 | 39 | goal ZF.thy "!!x A. [| 0 ~: A; x: A |] ==> EX y. y:x"; | 
| 40 | by (subgoal_tac "x ~= 0" 1); | |
| 3016 | 41 | by (ALLGOALS Blast_tac); | 
| 760 | 42 | qed "non_empty_family"; | 
| 484 | 43 | |
| 5137 | 44 | Goal "0 ~: A ==> EX f: A->Union(A). ALL x:A. f`x : x"; | 
| 484 | 45 | by (rtac AC_func 1); | 
| 46 | by (REPEAT (ares_tac [non_empty_family] 1)); | |
| 760 | 47 | qed "AC_func0"; | 
| 484 | 48 | |
| 5067 | 49 | Goal "EX f: (Pow(C)-{0}) -> C. ALL x:(Pow(C)-{0}). f`x : x";
 | 
| 484 | 50 | by (resolve_tac [AC_func0 RS bexE] 1); | 
| 51 | by (rtac bexI 2); | |
| 52 | by (assume_tac 2); | |
| 1461 | 53 | by (etac fun_weaken_type 2); | 
| 3016 | 54 | by (ALLGOALS Blast_tac); | 
| 760 | 55 | qed "AC_func_Pow"; | 
| 484 | 56 | |
| 5137 | 57 | Goal "0 ~: A ==> EX f. f: (PROD x:A. x)"; | 
| 1074 
d60f203eeddf
Modified proofs for new claset primitives.  The problem is that they enforce
 lcp parents: 
760diff
changeset | 58 | by (rtac AC_Pi 1); | 
| 
d60f203eeddf
Modified proofs for new claset primitives.  The problem is that they enforce
 lcp parents: 
760diff
changeset | 59 | 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 | 60 | qed "AC_Pi0"; | 
| 
d60f203eeddf
Modified proofs for new claset primitives.  The problem is that they enforce
 lcp parents: 
760diff
changeset | 61 | |
| 5321 | 62 | (*Used in Zorn.ML*) | 
| 63 | Goal "[| ch : (PROD X:Pow(S) - {0}. X);  X<=S;  X~=S |] ==> ch ` (S-X) : S-X";
 | |
| 64 | by (etac apply_type 1); | |
| 65 | by (blast_tac (claset() addSEs [equalityE]) 1); | |
| 66 | qed "choice_Diff"; | |
| 67 |