src/ZF/AC/AC2_AC6.ML
changeset 5068 fb28eaa07e01
parent 4091 771b1f6422a8
child 5137 60205b0de9b9
equal deleted inserted replaced
5067:62b6288e6005 5068:fb28eaa07e01
    12 
    12 
    13 (* ********************************************************************** *)
    13 (* ********************************************************************** *)
    14 (* AC1 ==> AC2                                                            *)
    14 (* AC1 ==> AC2                                                            *)
    15 (* ********************************************************************** *)
    15 (* ********************************************************************** *)
    16 
    16 
    17 goal thy "!!B. [| B:A; f:(PROD X:A. X); 0~:A |]  \
    17 Goal "!!B. [| B:A; f:(PROD X:A. X); 0~:A |]  \
    18 \               ==> {f`B} <= B Int {f`C. C:A}";
    18 \               ==> {f`B} <= B Int {f`C. C:A}";
    19 by (fast_tac (claset() addSEs [apply_type]) 1);
    19 by (fast_tac (claset() addSEs [apply_type]) 1);
    20 val lemma1 = result();
    20 val lemma1 = result();
    21 
    21 
    22 goalw thy [pairwise_disjoint_def]
    22 Goalw [pairwise_disjoint_def]
    23         "!!A. [| pairwise_disjoint(A); B:A; C:A; D:B; D:C |] ==> f`B = f`C";
    23         "!!A. [| pairwise_disjoint(A); B:A; C:A; D:B; D:C |] ==> f`B = f`C";
    24 by (fast_tac (claset() addSEs [equals0D]) 1);
    24 by (fast_tac (claset() addSEs [equals0D]) 1);
    25 val lemma2 = result();
    25 val lemma2 = result();
    26 
    26 
    27 goalw thy AC_defs "!!Z. AC1 ==> AC2"; 
    27 Goalw AC_defs "!!Z. AC1 ==> AC2"; 
    28 by (rtac allI 1);
    28 by (rtac allI 1);
    29 by (rtac impI 1);
    29 by (rtac impI 1);
    30 by (REPEAT (eresolve_tac [asm_rl,conjE,allE,exE,impE] 1));
    30 by (REPEAT (eresolve_tac [asm_rl,conjE,allE,exE,impE] 1));
    31 by (REPEAT (resolve_tac [exI,ballI,equalityI] 1));
    31 by (REPEAT (resolve_tac [exI,ballI,equalityI] 1));
    32 by (rtac lemma1 2 THEN (REPEAT (assume_tac 2)));
    32 by (rtac lemma1 2 THEN (REPEAT (assume_tac 2)));
    36 
    36 
    37 (* ********************************************************************** *)
    37 (* ********************************************************************** *)
    38 (* AC2 ==> AC1                                                            *)
    38 (* AC2 ==> AC1                                                            *)
    39 (* ********************************************************************** *)
    39 (* ********************************************************************** *)
    40 
    40 
    41 goal thy "!!A. 0~:A ==> 0 ~: {B*{B}. B:A}";
    41 Goal "!!A. 0~:A ==> 0 ~: {B*{B}. B:A}";
    42 by (fast_tac (claset() addSDs [sym RS (Sigma_empty_iff RS iffD1)]
    42 by (fast_tac (claset() addSDs [sym RS (Sigma_empty_iff RS iffD1)]
    43         addSEs [RepFunE, equals0D]) 1);
    43         addSEs [RepFunE, equals0D]) 1);
    44 val lemma1 = result();
    44 val lemma1 = result();
    45 
    45 
    46 goal thy "!!A. [| X*{X} Int C = {y}; X:A |]  \
    46 Goal "!!A. [| X*{X} Int C = {y}; X:A |]  \
    47 \               ==> (THE y. X*{X} Int C = {y}): X*A";
    47 \               ==> (THE y. X*{X} Int C = {y}): X*A";
    48 by (rtac subst_elem 1);
    48 by (rtac subst_elem 1);
    49 by (fast_tac (claset() addSIs [the_equality]
    49 by (fast_tac (claset() addSIs [the_equality]
    50                 addSEs [sym RS trans RS (singleton_eq_iff RS iffD1)]) 2);
    50                 addSEs [sym RS trans RS (singleton_eq_iff RS iffD1)]) 2);
    51 by (fast_tac (claset() addSEs [equalityE, make_elim singleton_subsetD]) 1);
    51 by (fast_tac (claset() addSEs [equalityE, make_elim singleton_subsetD]) 1);
    52 val lemma2 = result();
    52 val lemma2 = result();
    53 
    53 
    54 goal thy "!!A. ALL D:{E*{E}. E:A}. EX y. D Int C = {y}  \
    54 Goal "!!A. ALL D:{E*{E}. E:A}. EX y. D Int C = {y}  \
    55 \       ==> (lam x:A. fst(THE z. (x*{x} Int C = {z}))) :  \
    55 \       ==> (lam x:A. fst(THE z. (x*{x} Int C = {z}))) :  \
    56 \               (PROD X:A. X) ";
    56 \               (PROD X:A. X) ";
    57 by (fast_tac (claset() addSEs [lemma2]
    57 by (fast_tac (claset() addSEs [lemma2]
    58                 addSIs [lam_type, RepFunI, fst_type]
    58                 addSIs [lam_type, RepFunI, fst_type]
    59                 addSDs [bspec]) 1);
    59                 addSDs [bspec]) 1);
    60 val lemma3 = result();
    60 val lemma3 = result();
    61 
    61 
    62 goalw thy (AC_defs@AC_aux_defs) "!!Z. AC2 ==> AC1";
    62 Goalw (AC_defs@AC_aux_defs) "!!Z. AC2 ==> AC1";
    63 by (REPEAT (resolve_tac [allI, impI] 1));
    63 by (REPEAT (resolve_tac [allI, impI] 1));
    64 by (REPEAT (eresolve_tac [allE, impE] 1));
    64 by (REPEAT (eresolve_tac [allE, impE] 1));
    65 by (fast_tac (claset() addSEs [lemma3]) 2);
    65 by (fast_tac (claset() addSEs [lemma3]) 2);
    66 by (fast_tac (claset() addSIs [lemma1, equals0I]) 1);
    66 by (fast_tac (claset() addSIs [lemma1, equals0I]) 1);
    67 qed "AC2_AC1";
    67 qed "AC2_AC1";
    69 
    69 
    70 (* ********************************************************************** *)
    70 (* ********************************************************************** *)
    71 (* AC1 ==> AC4                                                            *)
    71 (* AC1 ==> AC4                                                            *)
    72 (* ********************************************************************** *)
    72 (* ********************************************************************** *)
    73 
    73 
    74 goal thy "!!R. 0 ~: {R``{x}. x:domain(R)}";
    74 Goal "!!R. 0 ~: {R``{x}. x:domain(R)}";
    75 by (fast_tac (claset() addEs [sym RS equals0D]) 1);
    75 by (fast_tac (claset() addEs [sym RS equals0D]) 1);
    76 val lemma = result();
    76 val lemma = result();
    77 
    77 
    78 goalw thy AC_defs "!!Z. AC1 ==> AC4";
    78 Goalw AC_defs "!!Z. AC1 ==> AC4";
    79 by (REPEAT (resolve_tac [allI, impI] 1));
    79 by (REPEAT (resolve_tac [allI, impI] 1));
    80 by (REPEAT (eresolve_tac [allE, lemma RSN (2, impE), exE] 1));
    80 by (REPEAT (eresolve_tac [allE, lemma RSN (2, impE), exE] 1));
    81 by (best_tac (claset() addSIs [lam_type] addSEs [apply_type]) 1);
    81 by (best_tac (claset() addSIs [lam_type] addSEs [apply_type]) 1);
    82 qed "AC1_AC4";
    82 qed "AC1_AC4";
    83 
    83 
    84 
    84 
    85 (* ********************************************************************** *)
    85 (* ********************************************************************** *)
    86 (* AC4 ==> AC3                                                            *)
    86 (* AC4 ==> AC3                                                            *)
    87 (* ********************************************************************** *)
    87 (* ********************************************************************** *)
    88 
    88 
    89 goal thy "!!f. f:A->B ==> (UN z:A. {z}*f`z) <= A*Union(B)";
    89 Goal "!!f. f:A->B ==> (UN z:A. {z}*f`z) <= A*Union(B)";
    90 by (fast_tac (claset() addSDs [apply_type]) 1);
    90 by (fast_tac (claset() addSDs [apply_type]) 1);
    91 val lemma1 = result();
    91 val lemma1 = result();
    92 
    92 
    93 goal thy "!!f. domain(UN z:A. {z}*f(z)) = {a:A. f(a)~=0}";
    93 Goal "!!f. domain(UN z:A. {z}*f(z)) = {a:A. f(a)~=0}";
    94 by (fast_tac (claset() addSIs [not_emptyI] addDs [range_type]) 1);
    94 by (fast_tac (claset() addSIs [not_emptyI] addDs [range_type]) 1);
    95 val lemma2 = result();
    95 val lemma2 = result();
    96 
    96 
    97 goal thy "!!f. x:A ==> (UN z:A. {z}*f(z))``{x} = f(x)";
    97 Goal "!!f. x:A ==> (UN z:A. {z}*f(z))``{x} = f(x)";
    98 by (Fast_tac 1);
    98 by (Fast_tac 1);
    99 val lemma3 = result();
    99 val lemma3 = result();
   100 
   100 
   101 goalw thy AC_defs "!!Z. AC4 ==> AC3";
   101 Goalw AC_defs "!!Z. AC4 ==> AC3";
   102 by (REPEAT (resolve_tac [allI,ballI] 1));
   102 by (REPEAT (resolve_tac [allI,ballI] 1));
   103 by (REPEAT (eresolve_tac [allE,impE] 1));
   103 by (REPEAT (eresolve_tac [allE,impE] 1));
   104 by (etac lemma1 1);
   104 by (etac lemma1 1);
   105 by (asm_full_simp_tac (simpset() addsimps [lemma2, lemma3]
   105 by (asm_full_simp_tac (simpset() addsimps [lemma2, lemma3]
   106                         addcongs [Pi_cong]) 1);
   106                         addcongs [Pi_cong]) 1);
   108 
   108 
   109 (* ********************************************************************** *)
   109 (* ********************************************************************** *)
   110 (* AC3 ==> AC1                                                            *)
   110 (* AC3 ==> AC1                                                            *)
   111 (* ********************************************************************** *)
   111 (* ********************************************************************** *)
   112 
   112 
   113 goal thy "!!A. b~:A ==> (PROD x:{a:A. id(A)`a~=b}. id(A)`x) = (PROD x:A. x)";
   113 Goal "!!A. b~:A ==> (PROD x:{a:A. id(A)`a~=b}. id(A)`x) = (PROD x:A. x)";
   114 by (asm_full_simp_tac (simpset() addsimps [id_def] addcongs [Pi_cong]) 1);
   114 by (asm_full_simp_tac (simpset() addsimps [id_def] addcongs [Pi_cong]) 1);
   115 by (res_inst_tac [("b","A")] subst_context 1);
   115 by (res_inst_tac [("b","A")] subst_context 1);
   116 by (Fast_tac 1);
   116 by (Fast_tac 1);
   117 val lemma = result();
   117 val lemma = result();
   118 
   118 
   119 goalw thy AC_defs "!!Z. AC3 ==> AC1";
   119 Goalw AC_defs "!!Z. AC3 ==> AC1";
   120 by (REPEAT (resolve_tac [allI, impI] 1));
   120 by (REPEAT (resolve_tac [allI, impI] 1));
   121 by (REPEAT (eresolve_tac [allE, ballE] 1));
   121 by (REPEAT (eresolve_tac [allE, ballE] 1));
   122 by (fast_tac (claset() addSIs [id_type]) 2);
   122 by (fast_tac (claset() addSIs [id_type]) 2);
   123 by (fast_tac (claset() addEs [lemma RS subst]) 1);
   123 by (fast_tac (claset() addEs [lemma RS subst]) 1);
   124 qed "AC3_AC1";
   124 qed "AC3_AC1";
   125 
   125 
   126 (* ********************************************************************** *)
   126 (* ********************************************************************** *)
   127 (* AC4 ==> AC5                                                            *)
   127 (* AC4 ==> AC5                                                            *)
   128 (* ********************************************************************** *)
   128 (* ********************************************************************** *)
   129 
   129 
   130 goalw thy (range_def::AC_defs) "!!Z. AC4 ==> AC5";
   130 Goalw (range_def::AC_defs) "!!Z. AC4 ==> AC5";
   131 by (REPEAT (resolve_tac [allI,ballI] 1));
   131 by (REPEAT (resolve_tac [allI,ballI] 1));
   132 by (REPEAT (eresolve_tac [allE,impE] 1));
   132 by (REPEAT (eresolve_tac [allE,impE] 1));
   133 by (eresolve_tac [fun_is_rel RS converse_type] 1);
   133 by (eresolve_tac [fun_is_rel RS converse_type] 1);
   134 by (etac exE 1);
   134 by (etac exE 1);
   135 by (rtac bexI 1);
   135 by (rtac bexI 1);
   146 
   146 
   147 (* ********************************************************************** *)
   147 (* ********************************************************************** *)
   148 (* AC5 ==> AC4, Rubin & Rubin, p. 11                                      *)
   148 (* AC5 ==> AC4, Rubin & Rubin, p. 11                                      *)
   149 (* ********************************************************************** *)
   149 (* ********************************************************************** *)
   150 
   150 
   151 goal thy "!!A. R <= A*B ==> (lam x:R. fst(x)) : R -> A";
   151 Goal "!!A. R <= A*B ==> (lam x:R. fst(x)) : R -> A";
   152 by (fast_tac (claset() addSIs [lam_type, fst_type]) 1);
   152 by (fast_tac (claset() addSIs [lam_type, fst_type]) 1);
   153 val lemma1 = result();
   153 val lemma1 = result();
   154 
   154 
   155 goalw thy [range_def] "!!A. R <= A*B ==> range(lam x:R. fst(x)) = domain(R)";
   155 Goalw [range_def] "!!A. R <= A*B ==> range(lam x:R. fst(x)) = domain(R)";
   156 by (rtac equalityI 1);
   156 by (rtac equalityI 1);
   157 by (fast_tac (claset() addSEs [lamE]
   157 by (fast_tac (claset() addSEs [lamE]
   158                 addEs [subst_elem]
   158                 addEs [subst_elem]
   159                 addSDs [Pair_fst_snd_eq]) 1);
   159                 addSDs [Pair_fst_snd_eq]) 1);
   160 by (rtac subsetI 1);
   160 by (rtac subsetI 1);
   161 by (etac domainE 1);
   161 by (etac domainE 1);
   162 by (rtac domainI 1);
   162 by (rtac domainI 1);
   163 by (fast_tac (claset() addSEs [lamI RS subst_elem] addIs [fst_conv RS ssubst]) 1);
   163 by (fast_tac (claset() addSEs [lamI RS subst_elem] addIs [fst_conv RS ssubst]) 1);
   164 val lemma2 = result();
   164 val lemma2 = result();
   165 
   165 
   166 goal thy "!!A. [| EX f: A->C. P(f,domain(f)); A=B |] ==>  EX f: B->C. P(f,B)";
   166 Goal "!!A. [| EX f: A->C. P(f,domain(f)); A=B |] ==>  EX f: B->C. P(f,B)";
   167 by (etac bexE 1);
   167 by (etac bexE 1);
   168 by (forward_tac [domain_of_fun] 1);
   168 by (forward_tac [domain_of_fun] 1);
   169 by (Fast_tac 1);
   169 by (Fast_tac 1);
   170 val lemma3 = result();
   170 val lemma3 = result();
   171 
   171 
   172 goal thy "!!g. [| R <= A*B; g: C->R; ALL x:C. (lam z:R. fst(z))` (g`x) = x |] \
   172 Goal "!!g. [| R <= A*B; g: C->R; ALL x:C. (lam z:R. fst(z))` (g`x) = x |] \
   173 \               ==> (lam x:C. snd(g`x)): (PROD x:C. R``{x})";
   173 \               ==> (lam x:C. snd(g`x)): (PROD x:C. R``{x})";
   174 by (rtac lam_type 1);
   174 by (rtac lam_type 1);
   175 by (dtac apply_type 1 THEN (assume_tac 1));
   175 by (dtac apply_type 1 THEN (assume_tac 1));
   176 by (dtac bspec 1 THEN (assume_tac 1));
   176 by (dtac bspec 1 THEN (assume_tac 1));
   177 by (rtac imageI 1);
   177 by (rtac imageI 1);
   178 by (resolve_tac [subsetD RS Pair_fst_snd_eq RSN (2, subst_elem)] 1
   178 by (resolve_tac [subsetD RS Pair_fst_snd_eq RSN (2, subst_elem)] 1
   179         THEN (REPEAT (assume_tac 1)));
   179         THEN (REPEAT (assume_tac 1)));
   180 by (Asm_full_simp_tac 1);
   180 by (Asm_full_simp_tac 1);
   181 val lemma4 = result();
   181 val lemma4 = result();
   182 
   182 
   183 goalw thy AC_defs "!!Z. AC5 ==> AC4";
   183 Goalw AC_defs "!!Z. AC5 ==> AC4";
   184 by (REPEAT (resolve_tac [allI,impI] 1));
   184 by (REPEAT (resolve_tac [allI,impI] 1));
   185 by (REPEAT (eresolve_tac [allE,ballE] 1));
   185 by (REPEAT (eresolve_tac [allE,ballE] 1));
   186 by (eresolve_tac [lemma1 RSN (2, notE)] 2 THEN (assume_tac 2));
   186 by (eresolve_tac [lemma1 RSN (2, notE)] 2 THEN (assume_tac 2));
   187 by (dresolve_tac [lemma2 RSN (2, lemma3)] 1 THEN (assume_tac 1));
   187 by (dresolve_tac [lemma2 RSN (2, lemma3)] 1 THEN (assume_tac 1));
   188 by (fast_tac (claset() addSEs [lemma4]) 1);
   188 by (fast_tac (claset() addSEs [lemma4]) 1);
   191 
   191 
   192 (* ********************************************************************** *)
   192 (* ********************************************************************** *)
   193 (* AC1 <-> AC6                                                            *)
   193 (* AC1 <-> AC6                                                            *)
   194 (* ********************************************************************** *)
   194 (* ********************************************************************** *)
   195 
   195 
   196 goalw thy AC_defs "AC1 <-> AC6";
   196 Goalw AC_defs "AC1 <-> AC6";
   197 by (fast_tac (claset() addDs [equals0D] addSEs [not_emptyE]) 1);
   197 by (fast_tac (claset() addDs [equals0D] addSEs [not_emptyE]) 1);
   198 qed "AC1_iff_AC6";
   198 qed "AC1_iff_AC6";
   199 
   199