src/ZF/AC.ML
author wenzelm
Fri, 08 Mar 2002 11:43:01 +0100
changeset 13046 69ab0e74ccda
parent 11320 56aa53caf333
permissions -rw-r--r--
removed Stanford mirror;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1074
diff changeset
     1
(*  Title:      ZF/AC.ML
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1074
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     5
5809
bacf85370ce0 removed obsolete comment and "open" declaration
paulson
parents: 5321
diff changeset
     6
The Axiom of Choice
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     7
*)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     8
11320
56aa53caf333 X-symbols for set theory
paulson
parents: 9907
diff changeset
     9
(*The same as AC, but no premise a \\<in> A*)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5809
diff changeset
    10
val [nonempty] = goal (the_context ())
11320
56aa53caf333 X-symbols for set theory
paulson
parents: 9907
diff changeset
    11
     "[| !!x. x \\<in> A ==> (\\<exists>y. y \\<in> B(x)) |] ==> \\<exists>z. z \\<in> Pi(A,B)";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    12
by (excluded_middle_tac "A=0" 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    13
by (asm_simp_tac (simpset() addsimps [Pi_empty1]) 2 THEN Blast_tac 2);
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    14
(*The non-trivial case*)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    15
by (blast_tac (claset() addIs [AC, nonempty]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 741
diff changeset
    16
qed "AC_Pi";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    17
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    18
(*Using dtac, this has the advantage of DELETING the universal quantifier*)
11320
56aa53caf333 X-symbols for set theory
paulson
parents: 9907
diff changeset
    19
Goal "\\<forall>x \\<in> A. \\<exists>y. y \\<in> B(x) ==> \\<exists>y. y \\<in> Pi(A,B)";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1074
diff changeset
    20
by (rtac AC_Pi 1);
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1074
diff changeset
    21
by (etac bspec 1);
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    22
by (assume_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 741
diff changeset
    23
qed "AC_ball_Pi";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    24
11320
56aa53caf333 X-symbols for set theory
paulson
parents: 9907
diff changeset
    25
Goal "\\<exists>f. f \\<in> (\\<Pi>X \\<in> Pow(C)-{0}. X)";
3840
e0baea4d485a fixed dots;
wenzelm
parents: 3016
diff changeset
    26
by (res_inst_tac [("B1", "%x. x")] (AC_Pi RS exE) 1);
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    27
by (etac exI 2);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2493
diff changeset
    28
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 741
diff changeset
    29
qed "AC_Pi_Pow";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    30
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5809
diff changeset
    31
val [nonempty] = goal (the_context ())
11320
56aa53caf333 X-symbols for set theory
paulson
parents: 9907
diff changeset
    32
     "[| !!x. x \\<in> A ==> (\\<exists>y. y \\<in> x)       \
56aa53caf333 X-symbols for set theory
paulson
parents: 9907
diff changeset
    33
\     |] ==> \\<exists>f \\<in> A->Union(A). \\<forall>x \\<in> A. f`x \\<in> x";
3840
e0baea4d485a fixed dots;
wenzelm
parents: 3016
diff changeset
    34
by (res_inst_tac [("B1", "%x. x")] (AC_Pi RS exE) 1);
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    35
by (etac nonempty 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    36
by (blast_tac (claset() addDs [apply_type] addIs [Pi_type]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 741
diff changeset
    37
qed "AC_func";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    38
11320
56aa53caf333 X-symbols for set theory
paulson
parents: 9907
diff changeset
    39
Goal "[| 0 \\<notin> A;  x \\<in> A |] ==> \\<exists>y. y \\<in> x";
56aa53caf333 X-symbols for set theory
paulson
parents: 9907
diff changeset
    40
by (subgoal_tac "x \\<noteq> 0" 1);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2493
diff changeset
    41
by (ALLGOALS Blast_tac);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 741
diff changeset
    42
qed "non_empty_family";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    43
11320
56aa53caf333 X-symbols for set theory
paulson
parents: 9907
diff changeset
    44
Goal "0 \\<notin> A ==> \\<exists>f \\<in> A->Union(A). \\<forall>x \\<in> A. f`x \\<in> x";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    45
by (rtac AC_func 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    46
by (REPEAT (ares_tac [non_empty_family] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 741
diff changeset
    47
qed "AC_func0";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    48
11320
56aa53caf333 X-symbols for set theory
paulson
parents: 9907
diff changeset
    49
Goal "\\<exists>f \\<in> (Pow(C)-{0}) -> C. \\<forall>x \\<in> Pow(C)-{0}. f`x \\<in> x";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    50
by (resolve_tac [AC_func0 RS bexE] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    51
by (rtac bexI 2);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    52
by (assume_tac 2);
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1074
diff changeset
    53
by (etac fun_weaken_type 2);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2493
diff changeset
    54
by (ALLGOALS Blast_tac);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 741
diff changeset
    55
qed "AC_func_Pow";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    56
11320
56aa53caf333 X-symbols for set theory
paulson
parents: 9907
diff changeset
    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
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    62
(*Used in Zorn.ML*)
11320
56aa53caf333 X-symbols for set theory
paulson
parents: 9907
diff changeset
    63
Goal "[| ch \\<in> (\\<Pi>X \\<in> Pow(S) - {0}. X);  X \\<subseteq> S;  X\\<noteq>S |] ==> ch ` (S-X) \\<in> S-X";
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    64
by (etac apply_type 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    65
by (blast_tac (claset() addSEs [equalityE]) 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    66
qed "choice_Diff";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    67