src/ZF/AC/WO_AC.ML
author paulson
Wed, 08 Jan 1997 15:04:27 +0100
changeset 2493 bdeb5024353a
parent 2469 b50b8c0eec01
child 3731 71366483323b
permissions -rw-r--r--
Removal of sum_cs and eq_cs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
     1
(*  Title:      ZF/AC/WO_AC.ML
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
     3
    Author:     Krzysztof Grabczewski
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     4
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     5
Lemmas used in the proofs like WO? ==> AC?
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     6
*)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     7
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     8
open WO_AC;
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     9
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    10
goal thy "!!A. [| well_ord(Union(A),r); 0~:A; B:A |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    11
\               ==> (THE b. first(b,B,r)) : B";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    12
by (fast_tac (!claset addSEs [well_ord_imp_ex1_first RS theI RS
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    13
                (first_def RS def_imp_iff RS iffD1 RS conjunct1)]) 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    14
val first_in_B = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    15
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    16
goal thy "!!A. [| well_ord(Union(A), R); 0~:A |] ==> EX f. f:(PROD X:A. X)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    17
by (fast_tac (!claset addSEs [first_in_B] addSIs [lam_type]) 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    18
val ex_choice_fun = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    19
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    20
goal thy "!!A. well_ord(A, R) ==> EX f. f:(PROD X: Pow(A)-{0}. X)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    21
by (fast_tac (!claset addSEs [well_ord_subset RS ex_choice_fun]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents: 1123
diff changeset
    22
val ex_choice_fun_Pow = result();