src/ZF/AC.ML
author lcp
Fri, 12 Aug 1994 12:51:34 +0200
changeset 516 1957113f0d7d
parent 484 70b789956bd3
child 741 5b0dedadb5c2
permissions -rw-r--r--
installation of new inductive/datatype sections
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     1
(*  Title: 	ZF/AC.ML
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     2
    ID:         $Id$
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
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
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     6
For AC.thy.  The Axiom of Choice
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     7
*)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     8
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     9
open AC;
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    10
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    11
(*The same as AC, but no premise a:A*)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    12
val [nonempty] = goal AC.thy
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    13
     "[| !!x. x:A ==> (EX y. y:B(x)) |] ==> EX z. z : Pi(A,B)";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    14
by (excluded_middle_tac "A=0" 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    15
by (asm_simp_tac (ZF_ss addsimps [Pi_empty1]) 2 THEN fast_tac ZF_cs 2);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    16
(*The non-trivial case*)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    17
by (safe_tac eq_cs);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    18
by (fast_tac (ZF_cs addSIs [AC, nonempty]) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    19
val AC_Pi = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    20
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    21
(*Using dtac, this has the advantage of DELETING the universal quantifier*)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    22
goal AC.thy "!!A B. ALL x:A. EX y. y:B(x) ==> EX y. y : Pi(A,B)";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    23
by (resolve_tac [AC_Pi] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    24
by (eresolve_tac [bspec] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    25
by (assume_tac 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    26
val AC_ball_Pi = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    27
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    28
goal AC.thy "EX f. f: (PROD X: Pow(C)-{0}. X)";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    29
by (res_inst_tac [("B1", "%x.x")] (AC_Pi RS exE) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    30
by (etac exI 2);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    31
by (fast_tac eq_cs 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    32
val AC_Pi_Pow = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    33
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    34
val [nonempty] = goal AC.thy
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    35
     "[| !!x. x:A ==> (EX y. y:x)	\
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    36
\     |] ==> EX f: A->Union(A). ALL x:A. f`x : x";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    37
by (res_inst_tac [("B1", "%x.x")] (AC_Pi RS exE) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    38
by (etac nonempty 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    39
by (fast_tac (ZF_cs addDs [apply_type] addIs [Pi_type]) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    40
val AC_func = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    41
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    42
goal AC.thy "!!x A. [| 0 ~: A;  x: A |] ==> EX y. y:x";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    43
by (resolve_tac [exCI] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    44
by (eresolve_tac [notE] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    45
by (resolve_tac [equals0I RS subst] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    46
by (eresolve_tac [spec RS notE] 1 THEN REPEAT (assume_tac 1));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    47
val non_empty_family = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    48
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    49
goal AC.thy "!!A. 0 ~: A ==> EX f: A->Union(A). ALL x:A. f`x : x";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    50
by (rtac AC_func 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    51
by (REPEAT (ares_tac [non_empty_family] 1));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    52
val AC_func0 = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    53
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    54
goal AC.thy "EX f: (Pow(C)-{0}) -> C. ALL x:(Pow(C)-{0}). f`x : x";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    55
by (resolve_tac [AC_func0 RS bexE] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    56
by (rtac bexI 2);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    57
by (assume_tac 2);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    58
by (eresolve_tac [fun_weaken_type] 2);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    59
by (ALLGOALS (fast_tac ZF_cs));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    60
val AC_func_Pow = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    61