src/ZF/AC/WO_AC.ML
author paulson
Fri, 03 Jan 1997 15:01:55 +0100
changeset 2469 b50b8c0eec01
parent 1461 6bcb44e4d6e5
child 3731 71366483323b
permissions -rw-r--r--
Implicit simpsets and clasets for FOL and ZF

(*  Title:      ZF/AC/WO_AC.ML
    ID:         $Id$
    Author:     Krzysztof Grabczewski

Lemmas used in the proofs like WO? ==> AC?
*)

open WO_AC;

goal thy "!!A. [| well_ord(Union(A),r); 0~:A; B:A |]  \
\               ==> (THE b. first(b,B,r)) : B";
by (fast_tac (!claset addSEs [well_ord_imp_ex1_first RS theI RS
                (first_def RS def_imp_iff RS iffD1 RS conjunct1)]) 1);
val first_in_B = result();

goal thy "!!A. [| well_ord(Union(A), R); 0~:A |] ==> EX f. f:(PROD X:A. X)";
by (fast_tac (!claset addSEs [first_in_B] addSIs [lam_type]) 1);
val ex_choice_fun = result();

goal thy "!!A. well_ord(A, R) ==> EX f. f:(PROD X: Pow(A)-{0}. X)";
by (fast_tac (!claset addSEs [well_ord_subset RS ex_choice_fun]) 1);
val ex_choice_fun_Pow = result();