src/ZF/AC.ML
changeset 741 5b0dedadb5c2
parent 484 70b789956bd3
child 760 f0200e91b272
equal deleted inserted replaced
740:281881c08397 741:5b0dedadb5c2
    12 val [nonempty] = goal AC.thy
    12 val [nonempty] = goal AC.thy
    13      "[| !!x. x:A ==> (EX y. y:B(x)) |] ==> EX z. z : Pi(A,B)";
    13      "[| !!x. x:A ==> (EX y. y:B(x)) |] ==> EX z. z : Pi(A,B)";
    14 by (excluded_middle_tac "A=0" 1);
    14 by (excluded_middle_tac "A=0" 1);
    15 by (asm_simp_tac (ZF_ss addsimps [Pi_empty1]) 2 THEN fast_tac ZF_cs 2);
    15 by (asm_simp_tac (ZF_ss addsimps [Pi_empty1]) 2 THEN fast_tac ZF_cs 2);
    16 (*The non-trivial case*)
    16 (*The non-trivial case*)
    17 by (safe_tac eq_cs);
    17 by (fast_tac (eq_cs addSIs [AC, nonempty]) 1);
    18 by (fast_tac (ZF_cs addSIs [AC, nonempty]) 1);
       
    19 val AC_Pi = result();
    18 val AC_Pi = result();
    20 
    19 
    21 (*Using dtac, this has the advantage of DELETING the universal quantifier*)
    20 (*Using dtac, this has the advantage of DELETING the universal quantifier*)
    22 goal AC.thy "!!A B. ALL x:A. EX y. y:B(x) ==> EX y. y : Pi(A,B)";
    21 goal AC.thy "!!A B. ALL x:A. EX y. y:B(x) ==> EX y. y : Pi(A,B)";
    23 by (resolve_tac [AC_Pi] 1);
    22 by (resolve_tac [AC_Pi] 1);
    37 by (res_inst_tac [("B1", "%x.x")] (AC_Pi RS exE) 1);
    36 by (res_inst_tac [("B1", "%x.x")] (AC_Pi RS exE) 1);
    38 by (etac nonempty 1);
    37 by (etac nonempty 1);
    39 by (fast_tac (ZF_cs addDs [apply_type] addIs [Pi_type]) 1);
    38 by (fast_tac (ZF_cs addDs [apply_type] addIs [Pi_type]) 1);
    40 val AC_func = result();
    39 val AC_func = result();
    41 
    40 
    42 goal AC.thy "!!x A. [| 0 ~: A;  x: A |] ==> EX y. y:x";
    41 goal ZF.thy "!!x A. [| 0 ~: A;  x: A |] ==> EX y. y:x";
    43 by (resolve_tac [exCI] 1);
    42 by (subgoal_tac "x ~= 0" 1);
    44 by (eresolve_tac [notE] 1);
    43 by (ALLGOALS (fast_tac eq_cs));
    45 by (resolve_tac [equals0I RS subst] 1);
       
    46 by (eresolve_tac [spec RS notE] 1 THEN REPEAT (assume_tac 1));
       
    47 val non_empty_family = result();
    44 val non_empty_family = result();
    48 
    45 
    49 goal AC.thy "!!A. 0 ~: A ==> EX f: A->Union(A). ALL x:A. f`x : x";
    46 goal AC.thy "!!A. 0 ~: A ==> EX f: A->Union(A). ALL x:A. f`x : x";
    50 by (rtac AC_func 1);
    47 by (rtac AC_func 1);
    51 by (REPEAT (ares_tac [non_empty_family] 1));
    48 by (REPEAT (ares_tac [non_empty_family] 1));