author | wenzelm |
Mon, 06 Oct 1997 19:15:22 +0200 | |
changeset 3795 | e687069e7257 |
parent 2877 | 6476784dba1c |
child 3840 | e0baea4d485a |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/upair |
0 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1991 University of Cambridge |
5 |
||
6 |
UNORDERED pairs in Zermelo-Fraenkel Set Theory |
|
7 |
||
8 |
Observe the order of dependence: |
|
9 |
Upair is defined in terms of Replace |
|
10 |
Un is defined in terms of Upair and Union (similarly for Int) |
|
11 |
cons is defined in terms of Upair and Un |
|
12 |
Ordered pairs and descriptions are defined using cons ("set notation") |
|
13 |
*) |
|
14 |
||
15 |
(*** Lemmas about power sets ***) |
|
16 |
||
1461 | 17 |
val Pow_bottom = empty_subsetI RS PowI; (* 0 : Pow(B) *) |
18 |
val Pow_top = subset_refl RS PowI; (* A : Pow(A) *) |
|
19 |
val Pow_neq_0 = Pow_top RSN (2,equals0D); (* Pow(a)=0 ==> P *) |
|
0 | 20 |
|
21 |
||
22 |
(*** Unordered pairs - Upair ***) |
|
23 |
||
2877 | 24 |
qed_goalw "Upair_iff" ZF.thy [Upair_def] |
0 | 25 |
"c : Upair(a,b) <-> (c=a | c=b)" |
2877 | 26 |
(fn _ => [ (blast_tac (!claset addEs [Pow_neq_0, sym RS Pow_neq_0]) 1) ]); |
2469 | 27 |
|
28 |
Addsimps [Upair_iff]; |
|
0 | 29 |
|
2877 | 30 |
qed_goal "UpairI1" ZF.thy "a : Upair(a,b)" |
2469 | 31 |
(fn _ => [ Simp_tac 1 ]); |
0 | 32 |
|
2877 | 33 |
qed_goal "UpairI2" ZF.thy "b : Upair(a,b)" |
2469 | 34 |
(fn _ => [ Simp_tac 1 ]); |
0 | 35 |
|
2877 | 36 |
qed_goal "UpairE" ZF.thy |
0 | 37 |
"[| a : Upair(b,c); a=b ==> P; a=c ==> P |] ==> P" |
38 |
(fn major::prems=> |
|
2469 | 39 |
[ (rtac (major RS (Upair_iff RS iffD1 RS disjE)) 1), |
0 | 40 |
(REPEAT (eresolve_tac prems 1)) ]); |
41 |
||
2469 | 42 |
AddSIs [UpairI1,UpairI2]; |
43 |
AddSEs [UpairE]; |
|
44 |
||
0 | 45 |
(*** Rules for binary union -- Un -- defined via Upair ***) |
46 |
||
2877 | 47 |
qed_goalw "Un_iff" ZF.thy [Un_def] "c : A Un B <-> (c:A | c:B)" |
48 |
(fn _ => [ Blast_tac 1 ]); |
|
2469 | 49 |
|
50 |
Addsimps [Un_iff]; |
|
0 | 51 |
|
2877 | 52 |
qed_goal "UnI1" ZF.thy "!!c. c : A ==> c : A Un B" |
2469 | 53 |
(fn _ => [ Asm_simp_tac 1 ]); |
0 | 54 |
|
2877 | 55 |
qed_goal "UnI2" ZF.thy "!!c. c : B ==> c : A Un B" |
2469 | 56 |
(fn _ => [ Asm_simp_tac 1 ]); |
57 |
||
2877 | 58 |
qed_goal "UnE" ZF.thy |
0 | 59 |
"[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P" |
60 |
(fn major::prems=> |
|
2469 | 61 |
[ (rtac (major RS (Un_iff RS iffD1 RS disjE)) 1), |
62 |
(REPEAT (eresolve_tac prems 1)) ]); |
|
0 | 63 |
|
572 | 64 |
(*Stronger version of the rule above*) |
2877 | 65 |
qed_goal "UnE'" ZF.thy |
572 | 66 |
"[| c : A Un B; c:A ==> P; [| c:B; c~:A |] ==> P |] ==> P" |
67 |
(fn major::prems => |
|
68 |
[(rtac (major RS UnE) 1), |
|
69 |
(eresolve_tac prems 1), |
|
70 |
(rtac classical 1), |
|
71 |
(eresolve_tac prems 1), |
|
72 |
(swap_res_tac prems 1), |
|
73 |
(etac notnotD 1)]); |
|
74 |
||
2469 | 75 |
(*Classical introduction rule: no commitment to A vs B*) |
2877 | 76 |
qed_goal "UnCI" ZF.thy "(c ~: B ==> c : A) ==> c : A Un B" |
2469 | 77 |
(fn prems=> |
2877 | 78 |
[ Simp_tac 1, blast_tac (!claset addSIs prems) 1 ]); |
0 | 79 |
|
2469 | 80 |
AddSIs [UnCI]; |
81 |
AddSEs [UnE]; |
|
0 | 82 |
|
83 |
||
84 |
(*** Rules for small intersection -- Int -- defined via Upair ***) |
|
85 |
||
2877 | 86 |
qed_goalw "Int_iff" ZF.thy [Int_def] "c : A Int B <-> (c:A & c:B)" |
87 |
(fn _ => [ Blast_tac 1 ]); |
|
2469 | 88 |
|
89 |
Addsimps [Int_iff]; |
|
90 |
||
2877 | 91 |
qed_goal "IntI" ZF.thy "!!c. [| c : A; c : B |] ==> c : A Int B" |
2469 | 92 |
(fn _ => [ Asm_simp_tac 1 ]); |
0 | 93 |
|
2877 | 94 |
qed_goal "IntD1" ZF.thy "!!c. c : A Int B ==> c : A" |
2469 | 95 |
(fn _ => [ Asm_full_simp_tac 1 ]); |
0 | 96 |
|
2877 | 97 |
qed_goal "IntD2" ZF.thy "!!c. c : A Int B ==> c : B" |
2469 | 98 |
(fn _ => [ Asm_full_simp_tac 1 ]); |
0 | 99 |
|
2877 | 100 |
qed_goal "IntE" ZF.thy |
0 | 101 |
"[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P" |
102 |
(fn prems=> |
|
103 |
[ (resolve_tac prems 1), |
|
104 |
(REPEAT (resolve_tac (prems RL [IntD1,IntD2]) 1)) ]); |
|
105 |
||
2469 | 106 |
AddSIs [IntI]; |
107 |
AddSEs [IntE]; |
|
0 | 108 |
|
109 |
(*** Rules for set difference -- defined via Upair ***) |
|
110 |
||
2877 | 111 |
qed_goalw "Diff_iff" ZF.thy [Diff_def] "c : A-B <-> (c:A & c~:B)" |
112 |
(fn _ => [ Blast_tac 1 ]); |
|
2469 | 113 |
|
114 |
Addsimps [Diff_iff]; |
|
115 |
||
2877 | 116 |
qed_goal "DiffI" ZF.thy "!!c. [| c : A; c ~: B |] ==> c : A - B" |
2469 | 117 |
(fn _ => [ Asm_simp_tac 1 ]); |
0 | 118 |
|
2877 | 119 |
qed_goal "DiffD1" ZF.thy "!!c. c : A - B ==> c : A" |
2469 | 120 |
(fn _ => [ Asm_full_simp_tac 1 ]); |
0 | 121 |
|
2877 | 122 |
qed_goal "DiffD2" ZF.thy "!!c. c : A - B ==> c ~: B" |
2469 | 123 |
(fn _ => [ Asm_full_simp_tac 1 ]); |
0 | 124 |
|
2877 | 125 |
qed_goal "DiffE" ZF.thy |
37 | 126 |
"[| c : A - B; [| c:A; c~:B |] ==> P |] ==> P" |
0 | 127 |
(fn prems=> |
128 |
[ (resolve_tac prems 1), |
|
485 | 129 |
(REPEAT (ares_tac (prems RL [DiffD1, DiffD2]) 1)) ]); |
0 | 130 |
|
2469 | 131 |
AddSIs [DiffI]; |
132 |
AddSEs [DiffE]; |
|
0 | 133 |
|
134 |
(*** Rules for cons -- defined via Un and Upair ***) |
|
135 |
||
2877 | 136 |
qed_goalw "cons_iff" ZF.thy [cons_def] "a : cons(b,A) <-> (a=b | a:A)" |
137 |
(fn _ => [ Blast_tac 1 ]); |
|
2469 | 138 |
|
139 |
Addsimps [cons_iff]; |
|
0 | 140 |
|
2877 | 141 |
qed_goal "consI1" ZF.thy "a : cons(a,B)" |
2469 | 142 |
(fn _ => [ Simp_tac 1 ]); |
143 |
||
144 |
Addsimps [consI1]; |
|
0 | 145 |
|
2877 | 146 |
qed_goal "consI2" ZF.thy "!!B. a : B ==> a : cons(b,B)" |
2469 | 147 |
(fn _ => [ Asm_simp_tac 1 ]); |
148 |
||
2877 | 149 |
qed_goal "consE" ZF.thy |
0 | 150 |
"[| a : cons(b,A); a=b ==> P; a:A ==> P |] ==> P" |
151 |
(fn major::prems=> |
|
2469 | 152 |
[ (rtac (major RS (cons_iff RS iffD1 RS disjE)) 1), |
0 | 153 |
(REPEAT (eresolve_tac (prems @ [UpairE]) 1)) ]); |
154 |
||
572 | 155 |
(*Stronger version of the rule above*) |
2877 | 156 |
qed_goal "consE'" ZF.thy |
572 | 157 |
"[| a : cons(b,A); a=b ==> P; [| a:A; a~=b |] ==> P |] ==> P" |
158 |
(fn major::prems => |
|
159 |
[(rtac (major RS consE) 1), |
|
160 |
(eresolve_tac prems 1), |
|
161 |
(rtac classical 1), |
|
162 |
(eresolve_tac prems 1), |
|
163 |
(swap_res_tac prems 1), |
|
164 |
(etac notnotD 1)]); |
|
165 |
||
2469 | 166 |
(*Classical introduction rule*) |
2877 | 167 |
qed_goal "consCI" ZF.thy "(a~:B ==> a=b) ==> a: cons(b,B)" |
2469 | 168 |
(fn prems=> |
2877 | 169 |
[ Simp_tac 1, blast_tac (!claset addSIs prems) 1 ]); |
0 | 170 |
|
2469 | 171 |
AddSIs [consCI]; |
172 |
AddSEs [consE]; |
|
0 | 173 |
|
2877 | 174 |
qed_goal "cons_not_0" ZF.thy "cons(a,B) ~= 0" |
175 |
(fn _ => [ (blast_tac (!claset addEs [equalityE]) 1) ]); |
|
1609 | 176 |
|
2469 | 177 |
bind_thm ("cons_neq_0", cons_not_0 RS notE); |
178 |
||
179 |
Addsimps [cons_not_0, cons_not_0 RS not_sym]; |
|
180 |
||
1609 | 181 |
|
0 | 182 |
(*** Singletons - using cons ***) |
183 |
||
2877 | 184 |
qed_goal "singleton_iff" ZF.thy "a : {b} <-> a=b" |
2469 | 185 |
(fn _ => [ Simp_tac 1 ]); |
186 |
||
2877 | 187 |
qed_goal "singletonI" ZF.thy "a : {a}" |
0 | 188 |
(fn _=> [ (rtac consI1 1) ]); |
189 |
||
2469 | 190 |
bind_thm ("singletonE", make_elim (singleton_iff RS iffD1)); |
0 | 191 |
|
2469 | 192 |
AddSIs [singletonI]; |
193 |
AddSEs [singletonE]; |
|
0 | 194 |
|
195 |
(*** Rules for Descriptions ***) |
|
196 |
||
2877 | 197 |
qed_goalw "the_equality" ZF.thy [the_def] |
0 | 198 |
"[| P(a); !!x. P(x) ==> x=a |] ==> (THE x. P(x)) = a" |
738 | 199 |
(fn [pa,eq] => |
2493 | 200 |
[ (fast_tac (!claset addSIs [pa] addEs [eq RS subst]) 1) ]); |
0 | 201 |
|
202 |
(* Only use this if you already know EX!x. P(x) *) |
|
2877 | 203 |
qed_goal "the_equality2" ZF.thy |
673 | 204 |
"!!P. [| EX! x. P(x); P(a) |] ==> (THE x. P(x)) = a" |
205 |
(fn _ => |
|
2469 | 206 |
[ (deepen_tac (!claset addSIs [the_equality]) 1 1) ]); |
0 | 207 |
|
2877 | 208 |
qed_goal "theI" ZF.thy "EX! x. P(x) ==> P(THE x. P(x))" |
0 | 209 |
(fn [major]=> |
210 |
[ (rtac (major RS ex1E) 1), |
|
211 |
(resolve_tac [major RS the_equality2 RS ssubst] 1), |
|
212 |
(REPEAT (assume_tac 1)) ]); |
|
213 |
||
686 | 214 |
(*Easier to apply than theI: conclusion has only one occurrence of P*) |
2877 | 215 |
qed_goal "theI2" ZF.thy |
686 | 216 |
"[| EX! x. P(x); !!x. P(x) ==> Q(x) |] ==> Q(THE x.P(x))" |
217 |
(fn prems => [ resolve_tac prems 1, |
|
1461 | 218 |
rtac theI 1, |
219 |
resolve_tac prems 1 ]); |
|
686 | 220 |
|
435 | 221 |
(*the_cong is no longer necessary: if (ALL y.P(y)<->Q(y)) then |
222 |
(THE x.P(x)) rewrites to (THE x. Q(x)) *) |
|
223 |
||
224 |
(*If it's "undefined", it's zero!*) |
|
2877 | 225 |
qed_goalw "the_0" ZF.thy [the_def] |
435 | 226 |
"!!P. ~ (EX! x. P(x)) ==> (THE x. P(x))=0" |
2877 | 227 |
(fn _ => [ (deepen_tac (!claset addSEs [ReplaceE]) 0 1) ]); |
435 | 228 |
|
0 | 229 |
|
230 |
(*** if -- a conditional expression for formulae ***) |
|
231 |
||
2877 | 232 |
goalw ZF.thy [if_def] "if(True,a,b) = a"; |
233 |
by (blast_tac (!claset addSIs [the_equality]) 1); |
|
760 | 234 |
qed "if_true"; |
0 | 235 |
|
2877 | 236 |
goalw ZF.thy [if_def] "if(False,a,b) = b"; |
237 |
by (blast_tac (!claset addSIs [the_equality]) 1); |
|
760 | 238 |
qed "if_false"; |
0 | 239 |
|
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
240 |
(*Never use with case splitting, or if P is known to be true or false*) |
2877 | 241 |
val prems = goalw ZF.thy [if_def] |
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
242 |
"[| P<->Q; Q ==> a=c; ~Q ==> b=d |] ==> if(P,a,b) = if(Q,c,d)"; |
2469 | 243 |
by (simp_tac (!simpset addsimps prems addcongs [conj_cong]) 1); |
760 | 244 |
qed "if_cong"; |
0 | 245 |
|
246 |
(*Not needed for rewriting, since P would rewrite to True anyway*) |
|
2877 | 247 |
goalw ZF.thy [if_def] "!!P. P ==> if(P,a,b) = a"; |
248 |
by (blast_tac (!claset addSIs [the_equality]) 1); |
|
760 | 249 |
qed "if_P"; |
0 | 250 |
|
251 |
(*Not needed for rewriting, since P would rewrite to False anyway*) |
|
2877 | 252 |
goalw ZF.thy [if_def] "!!P. ~P ==> if(P,a,b) = b"; |
253 |
by (blast_tac (!claset addSIs [the_equality]) 1); |
|
760 | 254 |
qed "if_not_P"; |
0 | 255 |
|
2469 | 256 |
Addsimps [if_true, if_false]; |
0 | 257 |
|
2877 | 258 |
qed_goal "expand_if" ZF.thy |
0 | 259 |
"P(if(Q,x,y)) <-> ((Q --> P(x)) & (~Q --> P(y)))" |
437 | 260 |
(fn _=> [ (excluded_middle_tac "Q" 1), |
2469 | 261 |
(Asm_simp_tac 1), |
262 |
(Asm_simp_tac 1) ]); |
|
0 | 263 |
|
2877 | 264 |
qed_goal "if_iff" ZF.thy "a: if(P,x,y) <-> P & a:x | ~P & a:y" |
2469 | 265 |
(fn _=> [ (simp_tac (!simpset setloop split_tac [expand_if]) 1) ]); |
1017
6a402dc505cf
Proved if_iff and used it to simplify proof of if_type.
lcp
parents:
985
diff
changeset
|
266 |
|
2877 | 267 |
qed_goal "if_type" ZF.thy |
1017
6a402dc505cf
Proved if_iff and used it to simplify proof of if_type.
lcp
parents:
985
diff
changeset
|
268 |
"[| P ==> a: A; ~P ==> b: A |] ==> if(P,a,b): A" |
6a402dc505cf
Proved if_iff and used it to simplify proof of if_type.
lcp
parents:
985
diff
changeset
|
269 |
(fn prems=> [ (simp_tac |
2469 | 270 |
(!simpset addsimps prems setloop split_tac [expand_if]) 1) ]); |
0 | 271 |
|
272 |
||
273 |
(*** Foundation lemmas ***) |
|
274 |
||
437 | 275 |
(*was called mem_anti_sym*) |
2877 | 276 |
qed_goal "mem_asym" ZF.thy "[| a:b; ~P ==> b:a |] ==> P" |
277 |
(fn prems=> |
|
278 |
[ (rtac classical 1), |
|
279 |
(res_inst_tac [("A1","{a,b}")] (foundation RS disjE) 1), |
|
280 |
REPEAT (blast_tac (!claset addIs prems addSEs [equalityE]) 1) ]); |
|
0 | 281 |
|
437 | 282 |
(*was called mem_anti_refl*) |
2877 | 283 |
qed_goal "mem_irrefl" ZF.thy "a:a ==> P" |
2469 | 284 |
(fn [major]=> [ (rtac ([major,major] MRS mem_asym) 1) ]); |
0 | 285 |
|
2469 | 286 |
(*mem_irrefl should NOT be added to default databases: |
287 |
it would be tried on most goals, making proofs slower!*) |
|
288 |
||
2877 | 289 |
qed_goal "mem_not_refl" ZF.thy "a ~: a" |
437 | 290 |
(K [ (rtac notI 1), (etac mem_irrefl 1) ]); |
0 | 291 |
|
435 | 292 |
(*Good for proving inequalities by rewriting*) |
2877 | 293 |
qed_goal "mem_imp_not_eq" ZF.thy "!!a A. a:A ==> a ~= A" |
294 |
(fn _=> [ blast_tac (!claset addSEs [mem_irrefl]) 1 ]); |
|
435 | 295 |
|
0 | 296 |
(*** Rules for succ ***) |
297 |
||
2877 | 298 |
qed_goalw "succ_iff" ZF.thy [succ_def] "i : succ(j) <-> i=j | i:j" |
299 |
(fn _ => [ Blast_tac 1 ]); |
|
2469 | 300 |
|
2877 | 301 |
qed_goalw "succI1" ZF.thy [succ_def] "i : succ(i)" |
0 | 302 |
(fn _=> [ (rtac consI1 1) ]); |
303 |
||
2469 | 304 |
Addsimps [succI1]; |
305 |
||
2877 | 306 |
qed_goalw "succI2" ZF.thy [succ_def] |
0 | 307 |
"i : j ==> i : succ(j)" |
308 |
(fn [prem]=> [ (rtac (prem RS consI2) 1) ]); |
|
309 |
||
2877 | 310 |
qed_goalw "succE" ZF.thy [succ_def] |
0 | 311 |
"[| i : succ(j); i=j ==> P; i:j ==> P |] ==> P" |
312 |
(fn major::prems=> |
|
313 |
[ (rtac (major RS consE) 1), |
|
314 |
(REPEAT (eresolve_tac prems 1)) ]); |
|
315 |
||
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
316 |
(*Classical introduction rule*) |
2877 | 317 |
qed_goal "succCI" ZF.thy "(i~:j ==> i=j) ==> i: succ(j)" |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
318 |
(fn [prem]=> |
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
319 |
[ (rtac (disjCI RS (succ_iff RS iffD2)) 1), |
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
320 |
(etac prem 1) ]); |
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
321 |
|
2469 | 322 |
AddSIs [succCI]; |
323 |
AddSEs [succE]; |
|
324 |
||
2877 | 325 |
qed_goal "succ_not_0" ZF.thy "succ(n) ~= 0" |
326 |
(fn _=> [ (blast_tac (!claset addSEs [equalityE]) 1) ]); |
|
0 | 327 |
|
2469 | 328 |
bind_thm ("succ_neq_0", succ_not_0 RS notE); |
329 |
||
330 |
Addsimps [succ_not_0, succ_not_0 RS not_sym]; |
|
331 |
AddSEs [succ_neq_0, sym RS succ_neq_0]; |
|
332 |
||
0 | 333 |
|
334 |
(* succ(c) <= B ==> c : B *) |
|
335 |
val succ_subsetD = succI1 RSN (2,subsetD); |
|
336 |
||
1609 | 337 |
(* succ(b) ~= b *) |
338 |
bind_thm ("succ_neq_self", succI1 RS mem_imp_not_eq RS not_sym); |
|
339 |
||
340 |
||
2877 | 341 |
qed_goal "succ_inject_iff" ZF.thy "succ(m) = succ(n) <-> m=n" |
342 |
(fn _=> [ (blast_tac (!claset addEs [mem_asym] addSEs [equalityE]) 1) ]); |
|
0 | 343 |
|
2469 | 344 |
bind_thm ("succ_inject", succ_inject_iff RS iffD1); |
0 | 345 |
|
2469 | 346 |
Addsimps [succ_inject_iff]; |
347 |
AddSDs [succ_inject]; |
|
0 | 348 |
|
2877 | 349 |
(*Not needed now that cons is available. Deletion reduces the search space.*) |
350 |
Delrules [UpairI1,UpairI2,UpairE]; |
|
2469 | 351 |
|
352 |
use"simpdata.ML"; |