src/ZF/AC/AC0_AC1.ML
author lcp
Tue, 25 Jul 1995 17:31:53 +0200
changeset 1196 d43c1f7a53fe
parent 1123 5dfdc1464966
child 1204 a4253da68be2
permissions -rw-r--r--
Numerous small improvements by KG and LCP

(*  Title: 	ZF/AC/AC0_AC1.ML
    ID:         $Id$
    Author: 	Krzysztof Gr`abczewski

AC0 is equivalent to AC1
AC0 comes from Suppes, AC1 from Rubin & Rubin
*)

goal thy "!!A. 0~:A ==> A <= Pow(Union(A))-{0}";
by (fast_tac AC_cs 1);
val subset_Pow_Union = result();

goal thy "!!f. [| f:(PROD X:A. X); D<=A |] ==> EX g. g:(PROD X:D. X)";
by (fast_tac (AC_cs addSIs [restrict_type, apply_type]) 1);
val lemma1 = result();

goalw thy AC_defs "!!Z. AC0 ==> AC1"; 
by (fast_tac (FOL_cs addSEs [lemma1, subset_Pow_Union]) 1);
qed "AC0_AC1";

goalw thy AC_defs "!!Z. AC1 ==> AC0";
by (fast_tac (FOL_cs addSIs [notI, singletonI] addSEs [notE, DiffE]) 1);
qed "AC1_AC0";