author | wenzelm |
Tue, 11 Dec 2001 16:00:26 +0100 | |
changeset 12466 | 5f4182667032 |
parent 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 |
||
11320 | 9 |
(*The same as AC, but no premise a \\<in> A*) |
9907 | 10 |
val [nonempty] = goal (the_context ()) |
11320 | 11 |
"[| !!x. x \\<in> A ==> (\\<exists>y. y \\<in> B(x)) |] ==> \\<exists>z. z \\<in> Pi(A,B)"; |
484 | 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*) |
|
11320 | 19 |
Goal "\\<forall>x \\<in> A. \\<exists>y. y \\<in> B(x) ==> \\<exists>y. y \\<in> 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 |
|
11320 | 25 |
Goal "\\<exists>f. f \\<in> (\\<Pi>X \\<in> 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 ()) |
11320 | 32 |
"[| !!x. x \\<in> A ==> (\\<exists>y. y \\<in> x) \ |
33 |
\ |] ==> \\<exists>f \\<in> A->Union(A). \\<forall>x \\<in> A. f`x \\<in> 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 |
|
11320 | 39 |
Goal "[| 0 \\<notin> A; x \\<in> A |] ==> \\<exists>y. y \\<in> x"; |
40 |
by (subgoal_tac "x \\<noteq> 0" 1); |
|
3016 | 41 |
by (ALLGOALS Blast_tac); |
760 | 42 |
qed "non_empty_family"; |
484 | 43 |
|
11320 | 44 |
Goal "0 \\<notin> A ==> \\<exists>f \\<in> A->Union(A). \\<forall>x \\<in> A. f`x \\<in> x"; |
484 | 45 |
by (rtac AC_func 1); |
46 |
by (REPEAT (ares_tac [non_empty_family] 1)); |
|
760 | 47 |
qed "AC_func0"; |
484 | 48 |
|
11320 | 49 |
Goal "\\<exists>f \\<in> (Pow(C)-{0}) -> C. \\<forall>x \\<in> Pow(C)-{0}. f`x \\<in> 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 |
|
11320 | 57 |
Goal "0 \\<notin> A ==> \\<exists>f. f \\<in> (\\<Pi>x \\<in> A. x)"; |
1074
d60f203eeddf
Modified proofs for new claset primitives. The problem is that they enforce
lcp
parents:
760
diff
changeset
|
58 |
by (rtac AC_Pi 1); |
d60f203eeddf
Modified proofs for new claset primitives. The problem is that they enforce
lcp
parents:
760
diff
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:
760
diff
changeset
|
60 |
qed "AC_Pi0"; |
d60f203eeddf
Modified proofs for new claset primitives. The problem is that they enforce
lcp
parents:
760
diff
changeset
|
61 |
|
5321 | 62 |
(*Used in Zorn.ML*) |
11320 | 63 |
Goal "[| ch \\<in> (\\<Pi>X \\<in> Pow(S) - {0}. X); X \\<subseteq> S; X\\<noteq>S |] ==> ch ` (S-X) \\<in> S-X"; |
5321 | 64 |
by (etac apply_type 1); |
65 |
by (blast_tac (claset() addSEs [equalityE]) 1); |
|
66 |
qed "choice_Diff"; |
|
67 |