| author | clasohm | 
| Tue, 24 Oct 1995 13:53:09 +0100 | |
| changeset 1291 | e173be970d27 | 
| parent 120 | 09287f26bfb8 | 
| permissions | -rw-r--r-- | 
| 0 | 1 | (* Title: ZF/zf.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson and Martin D Coen, CU Computer Laboratory | |
| 4 | Copyright 1992 University of Cambridge | |
| 5 | ||
| 6 | Basic introduction and elimination rules for Zermelo-Fraenkel Set Theory | |
| 7 | *) | |
| 8 | ||
| 9 | open ZF; | |
| 10 | ||
| 11 | signature ZF_LEMMAS = | |
| 12 | sig | |
| 13 | val ballE : thm | |
| 14 | val ballI : thm | |
| 15 | val ball_cong : thm | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 16 | val ball_simp : thm | 
| 0 | 17 | val ball_tac : int -> tactic | 
| 18 | val bexCI : thm | |
| 19 | val bexE : thm | |
| 20 | val bexI : thm | |
| 21 | val bex_cong : thm | |
| 22 | val bspec : thm | |
| 23 | val CollectD1 : thm | |
| 24 | val CollectD2 : thm | |
| 25 | val CollectE : thm | |
| 26 | val CollectI : thm | |
| 27 | val Collect_cong : thm | |
| 28 | val emptyE : thm | |
| 29 | val empty_subsetI : thm | |
| 30 | val equalityCE : thm | |
| 31 | val equalityD1 : thm | |
| 32 | val equalityD2 : thm | |
| 33 | val equalityE : thm | |
| 34 | val equalityI : thm | |
| 35 | val equality_iffI : thm | |
| 36 | val equals0D : thm | |
| 37 | val equals0I : thm | |
| 38 | val ex1_functional : thm | |
| 39 | val InterD : thm | |
| 40 | val InterE : thm | |
| 41 | val InterI : thm | |
| 42 | val INT_E : thm | |
| 43 | val INT_I : thm | |
| 44 | val lemmas_cs : claset | |
| 45 | val PowD : thm | |
| 46 | val PowI : thm | |
| 47 | val RepFunE : thm | |
| 48 | val RepFunI : thm | |
| 49 | val RepFun_eqI : thm | |
| 50 | val RepFun_cong : thm | |
| 51 | val ReplaceE : thm | |
| 52 | val ReplaceI : thm | |
| 53 | val Replace_iff : thm | |
| 54 | val Replace_cong : thm | |
| 55 | val rev_ballE : thm | |
| 56 | val rev_bspec : thm | |
| 57 | val rev_subsetD : thm | |
| 58 | val separation : thm | |
| 59 | val setup_induction : thm | |
| 60 | val set_mp_tac : int -> tactic | |
| 61 | val subsetCE : thm | |
| 62 | val subsetD : thm | |
| 63 | val subsetI : thm | |
| 64 | val subset_refl : thm | |
| 65 | val subset_trans : thm | |
| 66 | val UnionE : thm | |
| 67 | val UnionI : thm | |
| 68 | val UN_E : thm | |
| 69 | val UN_I : thm | |
| 70 | end; | |
| 71 | ||
| 72 | ||
| 73 | structure ZF_Lemmas : ZF_LEMMAS = | |
| 74 | struct | |
| 75 | ||
| 76 | (*** Bounded universal quantifier ***) | |
| 77 | ||
| 78 | val ballI = prove_goalw ZF.thy [Ball_def] | |
| 79 | "[| !!x. x:A ==> P(x) |] ==> ALL x:A. P(x)" | |
| 80 | (fn prems=> [ (REPEAT (ares_tac (prems @ [allI,impI]) 1)) ]); | |
| 81 | ||
| 82 | val bspec = prove_goalw ZF.thy [Ball_def] | |
| 83 | "[| ALL x:A. P(x); x: A |] ==> P(x)" | |
| 84 | (fn major::prems=> | |
| 85 | [ (rtac (major RS spec RS mp) 1), | |
| 86 | (resolve_tac prems 1) ]); | |
| 87 | ||
| 88 | val ballE = prove_goalw ZF.thy [Ball_def] | |
| 37 | 89 | "[| ALL x:A. P(x); P(x) ==> Q; x~:A ==> Q |] ==> Q" | 
| 0 | 90 | (fn major::prems=> | 
| 91 | [ (rtac (major RS allE) 1), | |
| 92 | (REPEAT (eresolve_tac (prems@[asm_rl,impCE]) 1)) ]); | |
| 93 | ||
| 94 | (*Used in the datatype package*) | |
| 95 | val rev_bspec = prove_goal ZF.thy | |
| 96 | "!!x A P. [| x: A; ALL x:A. P(x) |] ==> P(x)" | |
| 97 | (fn _ => | |
| 98 | [ REPEAT (ares_tac [bspec] 1) ]); | |
| 99 | ||
| 100 | (*Instantiates x first: better for automatic theorem proving?*) | |
| 101 | val rev_ballE = prove_goal ZF.thy | |
| 37 | 102 | "[| ALL x:A. P(x); x~:A ==> Q; P(x) ==> Q |] ==> Q" | 
| 0 | 103 | (fn major::prems=> | 
| 104 | [ (rtac (major RS ballE) 1), | |
| 105 | (REPEAT (eresolve_tac prems 1)) ]); | |
| 106 | ||
| 107 | (*Takes assumptions ALL x:A.P(x) and a:A; creates assumption P(a)*) | |
| 108 | val ball_tac = dtac bspec THEN' assume_tac; | |
| 109 | ||
| 110 | (*Trival rewrite rule; (ALL x:A.P)<->P holds only if A is nonempty!*) | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 111 | val ball_simp = prove_goal ZF.thy "(ALL x:A. True) <-> True" | 
| 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 112 | (fn _=> [ (REPEAT (ares_tac [TrueI,ballI,iffI] 1)) ]); | 
| 0 | 113 | |
| 114 | (*Congruence rule for rewriting*) | |
| 115 | val ball_cong = prove_goalw ZF.thy [Ball_def] | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 116 | "[| A=A'; !!x. x:A' ==> P(x) <-> P'(x) |] ==> Ball(A,P) <-> Ball(A',P')" | 
| 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 117 | (fn prems=> [ (simp_tac (FOL_ss addsimps prems) 1) ]); | 
| 0 | 118 | |
| 119 | (*** Bounded existential quantifier ***) | |
| 120 | ||
| 121 | val bexI = prove_goalw ZF.thy [Bex_def] | |
| 122 | "[| P(x); x: A |] ==> EX x:A. P(x)" | |
| 123 | (fn prems=> [ (REPEAT (ares_tac (prems @ [exI,conjI]) 1)) ]); | |
| 124 | ||
| 125 | (*Not of the general form for such rules; ~EX has become ALL~ *) | |
| 126 | val bexCI = prove_goal ZF.thy | |
| 127 | "[| ALL x:A. ~P(x) ==> P(a); a: A |] ==> EX x:A.P(x)" | |
| 128 | (fn prems=> | |
| 129 | [ (rtac classical 1), | |
| 130 | (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ]); | |
| 131 | ||
| 132 | val bexE = prove_goalw ZF.thy [Bex_def] | |
| 133 | "[| EX x:A. P(x); !!x. [| x:A; P(x) |] ==> Q \ | |
| 134 | \ |] ==> Q" | |
| 135 | (fn major::prems=> | |
| 136 | [ (rtac (major RS exE) 1), | |
| 137 | (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)) ]); | |
| 138 | ||
| 139 | (*We do not even have (EX x:A. True) <-> True unless A is nonempty!!*) | |
| 140 | ||
| 141 | val bex_cong = prove_goalw ZF.thy [Bex_def] | |
| 142 | "[| A=A'; !!x. x:A' ==> P(x) <-> P'(x) \ | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 143 | \ |] ==> Bex(A,P) <-> Bex(A',P')" | 
| 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 144 | (fn prems=> [ (simp_tac (FOL_ss addsimps prems addcongs [conj_cong]) 1) ]); | 
| 0 | 145 | |
| 146 | (*** Rules for subsets ***) | |
| 147 | ||
| 148 | val subsetI = prove_goalw ZF.thy [subset_def] | |
| 149 | "(!!x.x:A ==> x:B) ==> A <= B" | |
| 150 | (fn prems=> [ (REPEAT (ares_tac (prems @ [ballI]) 1)) ]); | |
| 151 | ||
| 152 | (*Rule in Modus Ponens style [was called subsetE] *) | |
| 153 | val subsetD = prove_goalw ZF.thy [subset_def] "[| A <= B; c:A |] ==> c:B" | |
| 154 | (fn major::prems=> | |
| 155 | [ (rtac (major RS bspec) 1), | |
| 156 | (resolve_tac prems 1) ]); | |
| 157 | ||
| 158 | (*Classical elimination rule*) | |
| 159 | val subsetCE = prove_goalw ZF.thy [subset_def] | |
| 37 | 160 | "[| A <= B; c~:A ==> P; c:B ==> P |] ==> P" | 
| 0 | 161 | (fn major::prems=> | 
| 162 | [ (rtac (major RS ballE) 1), | |
| 163 | (REPEAT (eresolve_tac prems 1)) ]); | |
| 164 | ||
| 165 | (*Takes assumptions A<=B; c:A and creates the assumption c:B *) | |
| 166 | val set_mp_tac = dtac subsetD THEN' assume_tac; | |
| 167 | ||
| 168 | (*Sometimes useful with premises in this order*) | |
| 169 | val rev_subsetD = prove_goal ZF.thy "!!A B c. [| c:A; A<=B |] ==> c:B" | |
| 170 | (fn _=> [REPEAT (ares_tac [subsetD] 1)]); | |
| 171 | ||
| 172 | val subset_refl = prove_goal ZF.thy "A <= A" | |
| 173 | (fn _=> [ (rtac subsetI 1), atac 1 ]); | |
| 174 | ||
| 175 | val subset_trans = prove_goal ZF.thy "[| A<=B; B<=C |] ==> A<=C" | |
| 176 | (fn prems=> [ (REPEAT (ares_tac ([subsetI]@(prems RL [subsetD])) 1)) ]); | |
| 177 | ||
| 178 | ||
| 179 | (*** Rules for equality ***) | |
| 180 | ||
| 181 | (*Anti-symmetry of the subset relation*) | |
| 182 | val equalityI = prove_goal ZF.thy "[| A <= B; B <= A |] ==> A = B" | |
| 183 | (fn prems=> [ (REPEAT (resolve_tac (prems@[conjI, extension RS iffD2]) 1)) ]); | |
| 184 | ||
| 185 | val equality_iffI = prove_goal ZF.thy "(!!x. x:A <-> x:B) ==> A = B" | |
| 186 | (fn [prem] => | |
| 187 | [ (rtac equalityI 1), | |
| 188 | (REPEAT (ares_tac [subsetI, prem RS iffD1, prem RS iffD2] 1)) ]); | |
| 189 | ||
| 190 | val equalityD1 = prove_goal ZF.thy "A = B ==> A<=B" | |
| 191 | (fn prems=> | |
| 192 | [ (rtac (extension RS iffD1 RS conjunct1) 1), | |
| 193 | (resolve_tac prems 1) ]); | |
| 194 | ||
| 195 | val equalityD2 = prove_goal ZF.thy "A = B ==> B<=A" | |
| 196 | (fn prems=> | |
| 197 | [ (rtac (extension RS iffD1 RS conjunct2) 1), | |
| 198 | (resolve_tac prems 1) ]); | |
| 199 | ||
| 200 | val equalityE = prove_goal ZF.thy | |
| 201 | "[| A = B; [| A<=B; B<=A |] ==> P |] ==> P" | |
| 202 | (fn prems=> | |
| 203 | [ (DEPTH_SOLVE (resolve_tac (prems@[equalityD1,equalityD2]) 1)) ]); | |
| 204 | ||
| 205 | val equalityCE = prove_goal ZF.thy | |
| 37 | 206 | "[| A = B; [| c:A; c:B |] ==> P; [| c~:A; c~:B |] ==> P |] ==> P" | 
| 0 | 207 | (fn major::prems=> | 
| 208 | [ (rtac (major RS equalityE) 1), | |
| 209 | (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1)) ]); | |
| 210 | ||
| 211 | (*Lemma for creating induction formulae -- for "pattern matching" on p | |
| 212 | To make the induction hypotheses usable, apply "spec" or "bspec" to | |
| 213 | put universal quantifiers over the free variables in p. | |
| 214 | Would it be better to do subgoal_tac "ALL z. p = f(z) --> R(z)" ??*) | |
| 215 | val setup_induction = prove_goal ZF.thy | |
| 216 | "[| p: A; !!z. z: A ==> p=z --> R |] ==> R" | |
| 217 | (fn prems=> | |
| 218 | [ (rtac mp 1), | |
| 219 | (REPEAT (resolve_tac (refl::prems) 1)) ]); | |
| 220 | ||
| 221 | ||
| 222 | (*** Rules for Replace -- the derived form of replacement ***) | |
| 223 | ||
| 224 | val ex1_functional = prove_goal ZF.thy | |
| 225 | "[| EX! z. P(a,z); P(a,b); P(a,c) |] ==> b = c" | |
| 226 | (fn prems=> | |
| 227 | [ (cut_facts_tac prems 1), | |
| 228 | (best_tac FOL_dup_cs 1) ]); | |
| 229 | ||
| 230 | val Replace_iff = prove_goalw ZF.thy [Replace_def] | |
| 231 |     "b : {y. x:A, P(x,y)}  <->  (EX x:A. P(x,b) & (ALL y. P(x,y) --> y=b))"
 | |
| 232 | (fn _=> | |
| 233 | [ (rtac (replacement RS iff_trans) 1), | |
| 234 | (REPEAT (ares_tac [refl,bex_cong,iffI,ballI,allI,conjI,impI,ex1I] 1 | |
| 235 | ORELSE eresolve_tac [conjE, spec RS mp, ex1_functional] 1)) ]); | |
| 236 | ||
| 237 | (*Introduction; there must be a unique y such that P(x,y), namely y=b. *) | |
| 238 | val ReplaceI = prove_goal ZF.thy | |
| 239 | "[| x: A; P(x,b); !!y. P(x,y) ==> y=b |] ==> \ | |
| 240 | \    b : {y. x:A, P(x,y)}"
 | |
| 241 | (fn prems=> | |
| 242 | [ (rtac (Replace_iff RS iffD2) 1), | |
| 243 | (REPEAT (ares_tac (prems@[bexI,conjI,allI,impI]) 1)) ]); | |
| 244 | ||
| 245 | (*Elimination; may asssume there is a unique y such that P(x,y), namely y=b. *) | |
| 246 | val ReplaceE = prove_goal ZF.thy | |
| 247 |     "[| b : {y. x:A, P(x,y)};  \
 | |
| 248 | \ !!x. [| x: A; P(x,b); ALL y. P(x,y)-->y=b |] ==> R \ | |
| 249 | \ |] ==> R" | |
| 250 | (fn prems=> | |
| 251 | [ (rtac (Replace_iff RS iffD1 RS bexE) 1), | |
| 252 | (etac conjE 2), | |
| 253 | (REPEAT (ares_tac prems 1)) ]); | |
| 254 | ||
| 255 | val Replace_cong = prove_goal ZF.thy | |
| 256 | "[| A=B; !!x y. x:B ==> P(x,y) <-> Q(x,y) |] ==> \ | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 257 | \ Replace(A,P) = Replace(B,Q)" | 
| 0 | 258 | (fn prems=> | 
| 259 | let val substprems = prems RL [subst, ssubst] | |
| 260 | and iffprems = prems RL [iffD1,iffD2] | |
| 261 | in [ (rtac equalityI 1), | |
| 262 | (REPEAT (eresolve_tac (substprems@[asm_rl, ReplaceE, spec RS mp]) 1 | |
| 263 | ORELSE resolve_tac [subsetI, ReplaceI] 1 | |
| 264 | ORELSE (resolve_tac iffprems 1 THEN assume_tac 2))) ] | |
| 265 | end); | |
| 266 | ||
| 267 | (*** Rules for RepFun ***) | |
| 268 | ||
| 269 | val RepFunI = prove_goalw ZF.thy [RepFun_def] | |
| 270 |     "!!a A. a : A ==> f(a) : {f(x). x:A}"
 | |
| 271 | (fn _ => [ (REPEAT (ares_tac [ReplaceI,refl] 1)) ]); | |
| 272 | ||
| 120 | 273 | (*Useful for coinduction proofs*) | 
| 0 | 274 | val RepFun_eqI = prove_goal ZF.thy | 
| 275 |     "!!b a f. [| b=f(a);  a : A |] ==> b : {f(x). x:A}"
 | |
| 276 | (fn _ => [ etac ssubst 1, etac RepFunI 1 ]); | |
| 277 | ||
| 278 | val RepFunE = prove_goalw ZF.thy [RepFun_def] | |
| 279 |     "[| b : {f(x). x:A};  \
 | |
| 280 | \ !!x.[| x:A; b=f(x) |] ==> P |] ==> \ | |
| 281 | \ P" | |
| 282 | (fn major::prems=> | |
| 283 | [ (rtac (major RS ReplaceE) 1), | |
| 284 | (REPEAT (ares_tac prems 1)) ]); | |
| 285 | ||
| 286 | val RepFun_cong = prove_goalw ZF.thy [RepFun_def] | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 287 | "[| A=B; !!x. x:B ==> f(x)=g(x) |] ==> RepFun(A,f) = RepFun(B,g)" | 
| 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 288 | (fn prems=> [ (simp_tac (FOL_ss addcongs [Replace_cong] addsimps prems) 1) ]); | 
| 0 | 289 | |
| 290 | ||
| 291 | (*** Rules for Collect -- forming a subset by separation ***) | |
| 292 | ||
| 293 | (*Separation is derivable from Replacement*) | |
| 294 | val separation = prove_goalw ZF.thy [Collect_def] | |
| 295 |     "a : {x:A. P(x)} <-> a:A & P(a)"
 | |
| 296 | (fn _=> [ (fast_tac (FOL_cs addIs [bexI,ReplaceI] | |
| 297 | addSEs [bexE,ReplaceE]) 1) ]); | |
| 298 | ||
| 299 | val CollectI = prove_goal ZF.thy | |
| 300 |     "[| a:A;  P(a) |] ==> a : {x:A. P(x)}"
 | |
| 301 | (fn prems=> | |
| 302 | [ (rtac (separation RS iffD2) 1), | |
| 303 | (REPEAT (resolve_tac (prems@[conjI]) 1)) ]); | |
| 304 | ||
| 305 | val CollectE = prove_goal ZF.thy | |
| 306 |     "[| a : {x:A. P(x)};  [| a:A; P(a) |] ==> R |] ==> R"
 | |
| 307 | (fn prems=> | |
| 308 | [ (rtac (separation RS iffD1 RS conjE) 1), | |
| 309 | (REPEAT (ares_tac prems 1)) ]); | |
| 310 | ||
| 311 | val CollectD1 = prove_goal ZF.thy "a : {x:A. P(x)} ==> a:A"
 | |
| 312 | (fn [major]=> | |
| 313 | [ (rtac (major RS CollectE) 1), | |
| 314 | (assume_tac 1) ]); | |
| 315 | ||
| 316 | val CollectD2 = prove_goal ZF.thy "a : {x:A. P(x)} ==> P(a)"
 | |
| 317 | (fn [major]=> | |
| 318 | [ (rtac (major RS CollectE) 1), | |
| 319 | (assume_tac 1) ]); | |
| 320 | ||
| 321 | val Collect_cong = prove_goalw ZF.thy [Collect_def] | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 322 | "[| A=B; !!x. x:B ==> P(x) <-> Q(x) |] ==> Collect(A,P) = Collect(B,Q)" | 
| 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 323 | (fn prems=> [ (simp_tac (FOL_ss addcongs [Replace_cong] addsimps prems) 1) ]); | 
| 0 | 324 | |
| 325 | (*** Rules for Unions ***) | |
| 326 | ||
| 327 | (*The order of the premises presupposes that C is rigid; A may be flexible*) | |
| 328 | val UnionI = prove_goal ZF.thy "[| B: C; A: B |] ==> A: Union(C)" | |
| 329 | (fn prems=> | |
| 330 | [ (resolve_tac [union_iff RS iffD2] 1), | |
| 331 | (REPEAT (resolve_tac (prems @ [bexI]) 1)) ]); | |
| 332 | ||
| 333 | val UnionE = prove_goal ZF.thy | |
| 334 | "[| A : Union(C); !!B.[| A: B; B: C |] ==> R |] ==> R" | |
| 335 | (fn prems=> | |
| 336 | [ (resolve_tac [union_iff RS iffD1 RS bexE] 1), | |
| 337 | (REPEAT (ares_tac prems 1)) ]); | |
| 338 | ||
| 339 | (*** Rules for Inter ***) | |
| 340 | ||
| 341 | (*Not obviously useful towards proving InterI, InterD, InterE*) | |
| 342 | val Inter_iff = prove_goalw ZF.thy [Inter_def,Ball_def] | |
| 343 | "A : Inter(C) <-> (ALL x:C. A: x) & (EX x. x:C)" | |
| 344 | (fn _=> [ (rtac (separation RS iff_trans) 1), | |
| 345 | (fast_tac (FOL_cs addIs [UnionI] addSEs [UnionE]) 1) ]); | |
| 346 | ||
| 347 | (* Intersection is well-behaved only if the family is non-empty! *) | |
| 348 | val InterI = prove_goalw ZF.thy [Inter_def] | |
| 349 | "[| !!x. x: C ==> A: x; c:C |] ==> A : Inter(C)" | |
| 350 | (fn prems=> | |
| 351 | [ (DEPTH_SOLVE (ares_tac ([CollectI,UnionI,ballI] @ prems) 1)) ]); | |
| 352 | ||
| 353 | (*A "destruct" rule -- every B in C contains A as an element, but | |
| 354 | A:B can hold when B:C does not! This rule is analogous to "spec". *) | |
| 355 | val InterD = prove_goalw ZF.thy [Inter_def] | |
| 356 | "[| A : Inter(C); B : C |] ==> A : B" | |
| 357 | (fn [major,minor]=> | |
| 358 | [ (rtac (major RS CollectD2 RS bspec) 1), | |
| 359 | (rtac minor 1) ]); | |
| 360 | ||
| 361 | (*"Classical" elimination rule -- does not require exhibiting B:C *) | |
| 362 | val InterE = prove_goalw ZF.thy [Inter_def] | |
| 37 | 363 | "[| A : Inter(C); A:B ==> R; B~:C ==> R |] ==> R" | 
| 0 | 364 | (fn major::prems=> | 
| 365 | [ (rtac (major RS CollectD2 RS ballE) 1), | |
| 366 | (REPEAT (eresolve_tac prems 1)) ]); | |
| 367 | ||
| 368 | (*** Rules for Unions of families ***) | |
| 369 | (* UN x:A. B(x) abbreviates Union({B(x). x:A}) *)
 | |
| 370 | ||
| 371 | (*The order of the premises presupposes that A is rigid; b may be flexible*) | |
| 372 | val UN_I = prove_goal ZF.thy "[| a: A; b: B(a) |] ==> b: (UN x:A. B(x))" | |
| 373 | (fn prems=> | |
| 374 | [ (REPEAT (resolve_tac (prems@[UnionI,RepFunI]) 1)) ]); | |
| 375 | ||
| 376 | val UN_E = prove_goal ZF.thy | |
| 377 | "[| b : (UN x:A. B(x)); !!x.[| x: A; b: B(x) |] ==> R |] ==> R" | |
| 378 | (fn major::prems=> | |
| 379 | [ (rtac (major RS UnionE) 1), | |
| 380 | (REPEAT (eresolve_tac (prems@[asm_rl, RepFunE, subst]) 1)) ]); | |
| 381 | ||
| 382 | ||
| 383 | (*** Rules for Intersections of families ***) | |
| 384 | (* INT x:A. B(x) abbreviates Inter({B(x). x:A}) *)
 | |
| 385 | ||
| 386 | val INT_I = prove_goal ZF.thy | |
| 387 | "[| !!x. x: A ==> b: B(x); a: A |] ==> b: (INT x:A. B(x))" | |
| 388 | (fn prems=> | |
| 389 | [ (REPEAT (ares_tac (prems@[InterI,RepFunI]) 1 | |
| 390 | ORELSE eresolve_tac [RepFunE,ssubst] 1)) ]); | |
| 391 | ||
| 392 | val INT_E = prove_goal ZF.thy | |
| 393 | "[| b : (INT x:A. B(x)); a: A |] ==> b : B(a)" | |
| 394 | (fn [major,minor]=> | |
| 395 | [ (rtac (major RS InterD) 1), | |
| 396 | (rtac (minor RS RepFunI) 1) ]); | |
| 397 | ||
| 398 | ||
| 399 | (*** Rules for Powersets ***) | |
| 400 | ||
| 401 | val PowI = prove_goal ZF.thy "A <= B ==> A : Pow(B)" | |
| 402 | (fn [prem]=> [ (rtac (prem RS (power_set RS iffD2)) 1) ]); | |
| 403 | ||
| 404 | val PowD = prove_goal ZF.thy "A : Pow(B) ==> A<=B" | |
| 405 | (fn [major]=> [ (rtac (major RS (power_set RS iffD1)) 1) ]); | |
| 406 | ||
| 407 | ||
| 408 | (*** Rules for the empty set ***) | |
| 409 | ||
| 410 | (*The set {x:0.False} is empty; by foundation it equals 0 
 | |
| 411 | See Suppes, page 21.*) | |
| 412 | val emptyE = prove_goal ZF.thy "a:0 ==> P" | |
| 413 | (fn [major]=> | |
| 414 | [ (rtac (foundation RS disjE) 1), | |
| 415 | (etac (equalityD2 RS subsetD RS CollectD2 RS FalseE) 1), | |
| 416 | (rtac major 1), | |
| 417 | (etac bexE 1), | |
| 418 | (etac (CollectD2 RS FalseE) 1) ]); | |
| 419 | ||
| 420 | val empty_subsetI = prove_goal ZF.thy "0 <= A" | |
| 421 | (fn _ => [ (REPEAT (ares_tac [equalityI,subsetI,emptyE] 1)) ]); | |
| 422 | ||
| 423 | val equals0I = prove_goal ZF.thy "[| !!y. y:A ==> False |] ==> A=0" | |
| 424 | (fn prems=> | |
| 425 | [ (REPEAT (ares_tac (prems@[empty_subsetI,subsetI,equalityI]) 1 | |
| 426 | ORELSE eresolve_tac (prems RL [FalseE]) 1)) ]); | |
| 427 | ||
| 428 | val equals0D = prove_goal ZF.thy "[| A=0; a:A |] ==> P" | |
| 429 | (fn [major,minor]=> | |
| 430 | [ (rtac (minor RS (major RS equalityD1 RS subsetD RS emptyE)) 1) ]); | |
| 431 | ||
| 432 | val lemmas_cs = FOL_cs | |
| 433 | addSIs [ballI, InterI, CollectI, PowI, subsetI] | |
| 434 | addIs [bexI, UnionI, ReplaceI, RepFunI] | |
| 435 | addSEs [bexE, make_elim PowD, UnionE, ReplaceE, RepFunE, | |
| 436 | CollectE, emptyE] | |
| 437 | addEs [rev_ballE, InterD, make_elim InterD, subsetD, subsetCE]; | |
| 438 | ||
| 439 | end; | |
| 440 | ||
| 441 | open ZF_Lemmas; |