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