| author | paulson | 
| Wed, 10 May 2000 11:17:01 +0200 | |
| changeset 8849 | f1933a670ae4 | 
| parent 7531 | 99c7e60d6b3f | 
| child 9180 | 3bda56c0d70d | 
| 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  | 
||
| 
825
 
76d9575950f2
Added Krzysztof's theorems subst_elem, not_emptyI, not_emptyE
 
lcp 
parents: 
775 
diff
changeset
 | 
9  | 
(*Useful examples: singletonI RS subst_elem, subst_elem RSN (2,IntI) *)  | 
| 5137 | 10  | 
Goal "[| b:A; a=b |] ==> a:A";  | 
| 
825
 
76d9575950f2
Added Krzysztof's theorems subst_elem, not_emptyI, not_emptyE
 
lcp 
parents: 
775 
diff
changeset
 | 
11  | 
by (etac ssubst 1);  | 
| 
 
76d9575950f2
Added Krzysztof's theorems subst_elem, not_emptyI, not_emptyE
 
lcp 
parents: 
775 
diff
changeset
 | 
12  | 
by (assume_tac 1);  | 
| 
 
76d9575950f2
Added Krzysztof's theorems subst_elem, not_emptyI, not_emptyE
 
lcp 
parents: 
775 
diff
changeset
 | 
13  | 
val subst_elem = result();  | 
| 
 
76d9575950f2
Added Krzysztof's theorems subst_elem, not_emptyI, not_emptyE
 
lcp 
parents: 
775 
diff
changeset
 | 
14  | 
|
| 2469 | 15  | 
|
| 0 | 16  | 
(*** Bounded universal quantifier ***)  | 
17  | 
||
| 775 | 18  | 
qed_goalw "ballI" ZF.thy [Ball_def]  | 
| 0 | 19  | 
"[| !!x. x:A ==> P(x) |] ==> ALL x:A. P(x)"  | 
20  | 
(fn prems=> [ (REPEAT (ares_tac (prems @ [allI,impI]) 1)) ]);  | 
|
21  | 
||
| 775 | 22  | 
qed_goalw "bspec" ZF.thy [Ball_def]  | 
| 0 | 23  | 
"[| ALL x:A. P(x); x: A |] ==> P(x)"  | 
24  | 
(fn major::prems=>  | 
|
25  | 
[ (rtac (major RS spec RS mp) 1),  | 
|
26  | 
(resolve_tac prems 1) ]);  | 
|
27  | 
||
| 775 | 28  | 
qed_goalw "ballE" ZF.thy [Ball_def]  | 
| 37 | 29  | 
"[| ALL x:A. P(x); P(x) ==> Q; x~:A ==> Q |] ==> Q"  | 
| 0 | 30  | 
(fn major::prems=>  | 
31  | 
[ (rtac (major RS allE) 1),  | 
|
32  | 
(REPEAT (eresolve_tac (prems@[asm_rl,impCE]) 1)) ]);  | 
|
33  | 
||
34  | 
(*Used in the datatype package*)  | 
|
| 775 | 35  | 
qed_goal "rev_bspec" ZF.thy  | 
| 0 | 36  | 
"!!x A P. [| x: A; ALL x:A. P(x) |] ==> P(x)"  | 
37  | 
(fn _ =>  | 
|
38  | 
[ REPEAT (ares_tac [bspec] 1) ]);  | 
|
39  | 
||
40  | 
(*Instantiates x first: better for automatic theorem proving?*)  | 
|
| 775 | 41  | 
qed_goal "rev_ballE" ZF.thy  | 
| 37 | 42  | 
"[| ALL x:A. P(x); x~:A ==> Q; P(x) ==> Q |] ==> Q"  | 
| 0 | 43  | 
(fn major::prems=>  | 
44  | 
[ (rtac (major RS ballE) 1),  | 
|
45  | 
(REPEAT (eresolve_tac prems 1)) ]);  | 
|
46  | 
||
| 2469 | 47  | 
AddSIs [ballI];  | 
48  | 
AddEs [rev_ballE];  | 
|
| 7531 | 49  | 
AddXDs [bspec];  | 
| 2469 | 50  | 
|
| 0 | 51  | 
(*Takes assumptions ALL x:A.P(x) and a:A; creates assumption P(a)*)  | 
52  | 
val ball_tac = dtac bspec THEN' assume_tac;  | 
|
53  | 
||
54  | 
(*Trival rewrite rule; (ALL x:A.P)<->P holds only if A is nonempty!*)  | 
|
| 3840 | 55  | 
qed_goal "ball_triv" ZF.thy "(ALL x:A. P) <-> ((EX x. x:A) --> P)"  | 
| 4091 | 56  | 
(fn _=> [ simp_tac (simpset() addsimps [Ball_def]) 1 ]);  | 
| 3425 | 57  | 
Addsimps [ball_triv];  | 
| 0 | 58  | 
|
59  | 
(*Congruence rule for rewriting*)  | 
|
| 775 | 60  | 
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
 | 
61  | 
"[| 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
 | 
62  | 
(fn prems=> [ (simp_tac (FOL_ss addsimps prems) 1) ]);  | 
| 0 | 63  | 
|
64  | 
(*** Bounded existential quantifier ***)  | 
|
65  | 
||
| 6287 | 66  | 
Goalw [Bex_def] "[| P(x); x: A |] ==> EX x:A. P(x)";  | 
67  | 
by (Blast_tac 1);  | 
|
68  | 
qed "bexI";  | 
|
69  | 
||
70  | 
(*The best argument order when there is only one x:A*)  | 
|
71  | 
Goalw [Bex_def] "[| x:A; P(x) |] ==> EX x:A. P(x)";  | 
|
72  | 
by (Blast_tac 1);  | 
|
73  | 
qed "rev_bexI";  | 
|
| 0 | 74  | 
|
75  | 
(*Not of the general form for such rules; ~EX has become ALL~ *)  | 
|
| 775 | 76  | 
qed_goal "bexCI" ZF.thy  | 
| 3840 | 77  | 
"[| ALL x:A. ~P(x) ==> P(a); a: A |] ==> EX x:A. P(x)"  | 
| 0 | 78  | 
(fn prems=>  | 
79  | 
[ (rtac classical 1),  | 
|
80  | 
(REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ]);  | 
|
81  | 
||
| 775 | 82  | 
qed_goalw "bexE" ZF.thy [Bex_def]  | 
| 0 | 83  | 
"[| EX x:A. P(x); !!x. [| x:A; P(x) |] ==> Q \  | 
84  | 
\ |] ==> Q"  | 
|
85  | 
(fn major::prems=>  | 
|
86  | 
[ (rtac (major RS exE) 1),  | 
|
87  | 
(REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)) ]);  | 
|
88  | 
||
| 2469 | 89  | 
AddIs [bexI];  | 
90  | 
AddSEs [bexE];  | 
|
91  | 
||
| 0 | 92  | 
(*We do not even have (EX x:A. True) <-> True unless A is nonempty!!*)  | 
| 3840 | 93  | 
qed_goal "bex_triv" ZF.thy "(EX x:A. P) <-> ((EX x. x:A) & P)"  | 
| 4091 | 94  | 
(fn _=> [ simp_tac (simpset() addsimps [Bex_def]) 1 ]);  | 
| 3425 | 95  | 
Addsimps [bex_triv];  | 
| 0 | 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  | 
|
| 2469 | 102  | 
Addcongs [ball_cong, bex_cong];  | 
103  | 
||
104  | 
||
| 0 | 105  | 
(*** Rules for subsets ***)  | 
106  | 
||
| 775 | 107  | 
qed_goalw "subsetI" ZF.thy [subset_def]  | 
| 3840 | 108  | 
"(!!x. x:A ==> x:B) ==> A <= B"  | 
| 0 | 109  | 
(fn prems=> [ (REPEAT (ares_tac (prems @ [ballI]) 1)) ]);  | 
110  | 
||
111  | 
(*Rule in Modus Ponens style [was called subsetE] *)  | 
|
| 775 | 112  | 
qed_goalw "subsetD" ZF.thy [subset_def] "[| A <= B; c:A |] ==> c:B"  | 
| 0 | 113  | 
(fn major::prems=>  | 
114  | 
[ (rtac (major RS bspec) 1),  | 
|
115  | 
(resolve_tac prems 1) ]);  | 
|
116  | 
||
117  | 
(*Classical elimination rule*)  | 
|
| 775 | 118  | 
qed_goalw "subsetCE" ZF.thy [subset_def]  | 
| 37 | 119  | 
"[| A <= B; c~:A ==> P; c:B ==> P |] ==> P"  | 
| 0 | 120  | 
(fn major::prems=>  | 
121  | 
[ (rtac (major RS ballE) 1),  | 
|
122  | 
(REPEAT (eresolve_tac prems 1)) ]);  | 
|
123  | 
||
| 2469 | 124  | 
AddSIs [subsetI];  | 
125  | 
AddEs [subsetCE, subsetD];  | 
|
126  | 
||
127  | 
||
| 0 | 128  | 
(*Takes assumptions A<=B; c:A and creates the assumption c:B *)  | 
129  | 
val set_mp_tac = dtac subsetD THEN' assume_tac;  | 
|
130  | 
||
131  | 
(*Sometimes useful with premises in this order*)  | 
|
| 775 | 132  | 
qed_goal "rev_subsetD" ZF.thy "!!A B c. [| c:A; A<=B |] ==> c:B"  | 
| 2877 | 133  | 
(fn _=> [ Blast_tac 1 ]);  | 
| 0 | 134  | 
|
| 6111 | 135  | 
(*Converts A<=B to x:A ==> x:B*)  | 
136  | 
fun impOfSubs th = th RSN (2, rev_subsetD);  | 
|
137  | 
||
| 1889 | 138  | 
qed_goal "contra_subsetD" ZF.thy "!!c. [| A <= B; c ~: B |] ==> c ~: A"  | 
| 2877 | 139  | 
(fn _=> [ Blast_tac 1 ]);  | 
| 1889 | 140  | 
|
141  | 
qed_goal "rev_contra_subsetD" ZF.thy "!!c. [| c ~: B; A <= B |] ==> c ~: A"  | 
|
| 2877 | 142  | 
(fn _=> [ Blast_tac 1 ]);  | 
| 1889 | 143  | 
|
| 775 | 144  | 
qed_goal "subset_refl" ZF.thy "A <= A"  | 
| 2877 | 145  | 
(fn _=> [ Blast_tac 1 ]);  | 
| 0 | 146  | 
|
| 2469 | 147  | 
Addsimps [subset_refl];  | 
148  | 
||
149  | 
qed_goal "subset_trans" ZF.thy "!!A B C. [| A<=B; B<=C |] ==> A<=C"  | 
|
| 2877 | 150  | 
(fn _=> [ Blast_tac 1 ]);  | 
| 0 | 151  | 
|
| 435 | 152  | 
(*Useful for proving A<=B by rewriting in some cases*)  | 
| 775 | 153  | 
qed_goalw "subset_iff" ZF.thy [subset_def,Ball_def]  | 
| 435 | 154  | 
"A<=B <-> (ALL x. x:A --> x:B)"  | 
155  | 
(fn _=> [ (rtac iff_refl 1) ]);  | 
|
156  | 
||
| 0 | 157  | 
|
158  | 
(*** Rules for equality ***)  | 
|
159  | 
||
160  | 
(*Anti-symmetry of the subset relation*)  | 
|
| 775 | 161  | 
qed_goal "equalityI" ZF.thy "[| A <= B; B <= A |] ==> A = B"  | 
| 0 | 162  | 
(fn prems=> [ (REPEAT (resolve_tac (prems@[conjI, extension RS iffD2]) 1)) ]);  | 
163  | 
||
| 2493 | 164  | 
AddIs [equalityI];  | 
165  | 
||
| 775 | 166  | 
qed_goal "equality_iffI" ZF.thy "(!!x. x:A <-> x:B) ==> A = B"  | 
| 0 | 167  | 
(fn [prem] =>  | 
168  | 
[ (rtac equalityI 1),  | 
|
169  | 
(REPEAT (ares_tac [subsetI, prem RS iffD1, prem RS iffD2] 1)) ]);  | 
|
170  | 
||
| 2469 | 171  | 
bind_thm ("equalityD1", extension RS iffD1 RS conjunct1);
 | 
172  | 
bind_thm ("equalityD2", extension RS iffD1 RS conjunct2);
 | 
|
| 0 | 173  | 
|
| 775 | 174  | 
qed_goal "equalityE" ZF.thy  | 
| 0 | 175  | 
"[| A = B; [| A<=B; B<=A |] ==> P |] ==> P"  | 
176  | 
(fn prems=>  | 
|
177  | 
[ (DEPTH_SOLVE (resolve_tac (prems@[equalityD1,equalityD2]) 1)) ]);  | 
|
178  | 
||
| 775 | 179  | 
qed_goal "equalityCE" ZF.thy  | 
| 37 | 180  | 
"[| A = B; [| c:A; c:B |] ==> P; [| c~:A; c~:B |] ==> P |] ==> P"  | 
| 0 | 181  | 
(fn major::prems=>  | 
182  | 
[ (rtac (major RS equalityE) 1),  | 
|
183  | 
(REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1)) ]);  | 
|
184  | 
||
185  | 
(*Lemma for creating induction formulae -- for "pattern matching" on p  | 
|
186  | 
To make the induction hypotheses usable, apply "spec" or "bspec" to  | 
|
187  | 
put universal quantifiers over the free variables in p.  | 
|
188  | 
Would it be better to do subgoal_tac "ALL z. p = f(z) --> R(z)" ??*)  | 
|
| 775 | 189  | 
qed_goal "setup_induction" ZF.thy  | 
| 0 | 190  | 
"[| p: A; !!z. z: A ==> p=z --> R |] ==> R"  | 
191  | 
(fn prems=>  | 
|
192  | 
[ (rtac mp 1),  | 
|
193  | 
(REPEAT (resolve_tac (refl::prems) 1)) ]);  | 
|
194  | 
||
195  | 
||
196  | 
(*** Rules for Replace -- the derived form of replacement ***)  | 
|
197  | 
||
| 775 | 198  | 
qed_goalw "Replace_iff" ZF.thy [Replace_def]  | 
| 0 | 199  | 
    "b : {y. x:A, P(x,y)}  <->  (EX x:A. P(x,b) & (ALL y. P(x,y) --> y=b))"
 | 
200  | 
(fn _=>  | 
|
201  | 
[ (rtac (replacement RS iff_trans) 1),  | 
|
202  | 
(REPEAT (ares_tac [refl,bex_cong,iffI,ballI,allI,conjI,impI,ex1I] 1  | 
|
203  | 
ORELSE eresolve_tac [conjE, spec RS mp, ex1_functional] 1)) ]);  | 
|
204  | 
||
205  | 
(*Introduction; there must be a unique y such that P(x,y), namely y=b. *)  | 
|
| 775 | 206  | 
qed_goal "ReplaceI" ZF.thy  | 
| 485 | 207  | 
"[| P(x,b); x: A; !!y. P(x,y) ==> y=b |] ==> \  | 
| 0 | 208  | 
\    b : {y. x:A, P(x,y)}"
 | 
209  | 
(fn prems=>  | 
|
210  | 
[ (rtac (Replace_iff RS iffD2) 1),  | 
|
211  | 
(REPEAT (ares_tac (prems@[bexI,conjI,allI,impI]) 1)) ]);  | 
|
212  | 
||
213  | 
(*Elimination; may asssume there is a unique y such that P(x,y), namely y=b. *)  | 
|
| 775 | 214  | 
qed_goal "ReplaceE" ZF.thy  | 
| 0 | 215  | 
    "[| b : {y. x:A, P(x,y)};  \
 | 
216  | 
\ !!x. [| x: A; P(x,b); ALL y. P(x,y)-->y=b |] ==> R \  | 
|
217  | 
\ |] ==> R"  | 
|
218  | 
(fn prems=>  | 
|
219  | 
[ (rtac (Replace_iff RS iffD1 RS bexE) 1),  | 
|
220  | 
(etac conjE 2),  | 
|
221  | 
(REPEAT (ares_tac prems 1)) ]);  | 
|
222  | 
||
| 485 | 223  | 
(*As above but without the (generally useless) 3rd assumption*)  | 
| 775 | 224  | 
qed_goal "ReplaceE2" ZF.thy  | 
| 485 | 225  | 
    "[| b : {y. x:A, P(x,y)};  \
 | 
226  | 
\ !!x. [| x: A; P(x,b) |] ==> R \  | 
|
227  | 
\ |] ==> R"  | 
|
228  | 
(fn major::prems=>  | 
|
229  | 
[ (rtac (major RS ReplaceE) 1),  | 
|
230  | 
(REPEAT (ares_tac prems 1)) ]);  | 
|
231  | 
||
| 2469 | 232  | 
AddIs [ReplaceI];  | 
233  | 
AddSEs [ReplaceE2];  | 
|
234  | 
||
| 775 | 235  | 
qed_goal "Replace_cong" ZF.thy  | 
| 0 | 236  | 
"[| 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
 | 
237  | 
\ Replace(A,P) = Replace(B,Q)"  | 
| 0 | 238  | 
(fn prems=>  | 
239  | 
let val substprems = prems RL [subst, ssubst]  | 
|
240  | 
and iffprems = prems RL [iffD1,iffD2]  | 
|
241  | 
in [ (rtac equalityI 1),  | 
|
| 1461 | 242  | 
(REPEAT (eresolve_tac (substprems@[asm_rl, ReplaceE, spec RS mp]) 1  | 
243  | 
ORELSE resolve_tac [subsetI, ReplaceI] 1  | 
|
244  | 
ORELSE (resolve_tac iffprems 1 THEN assume_tac 2))) ]  | 
|
| 0 | 245  | 
end);  | 
246  | 
||
| 2469 | 247  | 
Addcongs [Replace_cong];  | 
248  | 
||
| 0 | 249  | 
(*** Rules for RepFun ***)  | 
250  | 
||
| 775 | 251  | 
qed_goalw "RepFunI" ZF.thy [RepFun_def]  | 
| 0 | 252  | 
    "!!a A. a : A ==> f(a) : {f(x). x:A}"
 | 
253  | 
(fn _ => [ (REPEAT (ares_tac [ReplaceI,refl] 1)) ]);  | 
|
254  | 
||
| 120 | 255  | 
(*Useful for coinduction proofs*)  | 
| 775 | 256  | 
qed_goal "RepFun_eqI" ZF.thy  | 
| 0 | 257  | 
    "!!b a f. [| b=f(a);  a : A |] ==> b : {f(x). x:A}"
 | 
258  | 
(fn _ => [ etac ssubst 1, etac RepFunI 1 ]);  | 
|
259  | 
||
| 775 | 260  | 
qed_goalw "RepFunE" ZF.thy [RepFun_def]  | 
| 0 | 261  | 
    "[| b : {f(x). x:A};  \
 | 
262  | 
\ !!x.[| x:A; b=f(x) |] ==> P |] ==> \  | 
|
263  | 
\ P"  | 
|
264  | 
(fn major::prems=>  | 
|
265  | 
[ (rtac (major RS ReplaceE) 1),  | 
|
266  | 
(REPEAT (ares_tac prems 1)) ]);  | 
|
267  | 
||
| 2716 | 268  | 
AddIs [RepFun_eqI];  | 
| 2469 | 269  | 
AddSEs [RepFunE];  | 
270  | 
||
| 775 | 271  | 
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
 | 
272  | 
"[| A=B; !!x. x:B ==> f(x)=g(x) |] ==> RepFun(A,f) = RepFun(B,g)"  | 
| 4091 | 273  | 
(fn prems=> [ (simp_tac (simpset() addsimps prems) 1) ]);  | 
| 2469 | 274  | 
|
275  | 
Addcongs [RepFun_cong];  | 
|
| 0 | 276  | 
|
| 775 | 277  | 
qed_goalw "RepFun_iff" ZF.thy [Bex_def]  | 
| 485 | 278  | 
    "b : {f(x). x:A} <-> (EX x:A. b=f(x))"
 | 
| 2877 | 279  | 
(fn _ => [Blast_tac 1]);  | 
| 485 | 280  | 
|
| 5067 | 281  | 
Goal "{x. x:A} = A";
 | 
| 2877 | 282  | 
by (Blast_tac 1);  | 
| 2469 | 283  | 
qed "triv_RepFun";  | 
284  | 
||
285  | 
Addsimps [RepFun_iff, triv_RepFun];  | 
|
| 0 | 286  | 
|
287  | 
(*** Rules for Collect -- forming a subset by separation ***)  | 
|
288  | 
||
289  | 
(*Separation is derivable from Replacement*)  | 
|
| 775 | 290  | 
qed_goalw "separation" ZF.thy [Collect_def]  | 
| 0 | 291  | 
    "a : {x:A. P(x)} <-> a:A & P(a)"
 | 
| 2877 | 292  | 
(fn _=> [Blast_tac 1]);  | 
| 2469 | 293  | 
|
294  | 
Addsimps [separation];  | 
|
| 0 | 295  | 
|
| 775 | 296  | 
qed_goal "CollectI" ZF.thy  | 
| 2469 | 297  | 
    "!!P. [| a:A;  P(a) |] ==> a : {x:A. P(x)}"
 | 
298  | 
(fn _=> [ Asm_simp_tac 1 ]);  | 
|
| 0 | 299  | 
|
| 775 | 300  | 
qed_goal "CollectE" ZF.thy  | 
| 0 | 301  | 
    "[| a : {x:A. P(x)};  [| a:A; P(a) |] ==> R |] ==> R"
 | 
302  | 
(fn prems=>  | 
|
303  | 
[ (rtac (separation RS iffD1 RS conjE) 1),  | 
|
304  | 
(REPEAT (ares_tac prems 1)) ]);  | 
|
305  | 
||
| 2469 | 306  | 
qed_goal "CollectD1" ZF.thy "!!P. a : {x:A. P(x)} ==> a:A"
 | 
307  | 
(fn _=> [ (etac CollectE 1), (assume_tac 1) ]);  | 
|
| 0 | 308  | 
|
| 2469 | 309  | 
qed_goal "CollectD2" ZF.thy "!!P. a : {x:A. P(x)} ==> P(a)"
 | 
310  | 
(fn _=> [ (etac CollectE 1), (assume_tac 1) ]);  | 
|
| 0 | 311  | 
|
| 775 | 312  | 
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
 | 
313  | 
"[| A=B; !!x. x:B ==> P(x) <-> Q(x) |] ==> Collect(A,P) = Collect(B,Q)"  | 
| 4091 | 314  | 
(fn prems=> [ (simp_tac (simpset() addsimps prems) 1) ]);  | 
| 2469 | 315  | 
|
316  | 
AddSIs [CollectI];  | 
|
317  | 
AddSEs [CollectE];  | 
|
318  | 
Addcongs [Collect_cong];  | 
|
| 0 | 319  | 
|
320  | 
(*** Rules for Unions ***)  | 
|
321  | 
||
| 2469 | 322  | 
Addsimps [Union_iff];  | 
323  | 
||
| 0 | 324  | 
(*The order of the premises presupposes that C is rigid; A may be flexible*)  | 
| 2469 | 325  | 
qed_goal "UnionI" ZF.thy "!!C. [| B: C; A: B |] ==> A: Union(C)"  | 
| 2877 | 326  | 
(fn _=> [ Simp_tac 1, Blast_tac 1 ]);  | 
| 0 | 327  | 
|
| 775 | 328  | 
qed_goal "UnionE" ZF.thy  | 
| 0 | 329  | 
"[| A : Union(C); !!B.[| A: B; B: C |] ==> R |] ==> R"  | 
330  | 
(fn prems=>  | 
|
| 485 | 331  | 
[ (resolve_tac [Union_iff RS iffD1 RS bexE] 1),  | 
| 0 | 332  | 
(REPEAT (ares_tac prems 1)) ]);  | 
333  | 
||
334  | 
(*** Rules for Unions of families ***)  | 
|
335  | 
(* UN x:A. B(x) abbreviates Union({B(x). x:A}) *)
 | 
|
336  | 
||
| 775 | 337  | 
qed_goalw "UN_iff" ZF.thy [Bex_def]  | 
| 485 | 338  | 
"b : (UN x:A. B(x)) <-> (EX x:A. b : B(x))"  | 
| 2877 | 339  | 
(fn _=> [ Simp_tac 1, Blast_tac 1 ]);  | 
| 2469 | 340  | 
|
341  | 
Addsimps [UN_iff];  | 
|
| 485 | 342  | 
|
| 0 | 343  | 
(*The order of the premises presupposes that A is rigid; b may be flexible*)  | 
| 2469 | 344  | 
qed_goal "UN_I" ZF.thy "!!A B. [| a: A; b: B(a) |] ==> b: (UN x:A. B(x))"  | 
| 2877 | 345  | 
(fn _=> [ Simp_tac 1, Blast_tac 1 ]);  | 
| 0 | 346  | 
|
| 775 | 347  | 
qed_goal "UN_E" ZF.thy  | 
| 0 | 348  | 
"[| b : (UN x:A. B(x)); !!x.[| x: A; b: B(x) |] ==> R |] ==> R"  | 
349  | 
(fn major::prems=>  | 
|
350  | 
[ (rtac (major RS UnionE) 1),  | 
|
351  | 
(REPEAT (eresolve_tac (prems@[asm_rl, RepFunE, subst]) 1)) ]);  | 
|
352  | 
||
| 775 | 353  | 
qed_goal "UN_cong" ZF.thy  | 
| 3840 | 354  | 
"[| A=B; !!x. x:B ==> C(x)=D(x) |] ==> (UN x:A. C(x)) = (UN x:B. D(x))"  | 
| 4091 | 355  | 
(fn prems=> [ (simp_tac (simpset() addsimps prems) 1) ]);  | 
| 2469 | 356  | 
|
357  | 
(*No "Addcongs [UN_cong]" because UN is a combination of constants*)  | 
|
358  | 
||
359  | 
(* UN_E appears before UnionE so that it is tried first, to avoid expensive  | 
|
360  | 
calls to hyp_subst_tac. Cannot include UN_I as it is unsafe: would enlarge  | 
|
361  | 
the search space.*)  | 
|
362  | 
AddIs [UnionI];  | 
|
363  | 
AddSEs [UN_E];  | 
|
364  | 
AddSEs [UnionE];  | 
|
365  | 
||
366  | 
||
367  | 
(*** Rules for Inter ***)  | 
|
368  | 
||
369  | 
(*Not obviously useful towards proving InterI, InterD, InterE*)  | 
|
370  | 
qed_goalw "Inter_iff" ZF.thy [Inter_def,Ball_def]  | 
|
371  | 
"A : Inter(C) <-> (ALL x:C. A: x) & (EX x. x:C)"  | 
|
| 2877 | 372  | 
(fn _=> [ Simp_tac 1, Blast_tac 1 ]);  | 
| 435 | 373  | 
|
| 2469 | 374  | 
(* Intersection is well-behaved only if the family is non-empty! *)  | 
| 
2815
 
c05fa3ce5439
Improved intersection rule InterI: now truly safe, since the unsafeness is
 
paulson 
parents: 
2716 
diff
changeset
 | 
375  | 
qed_goal "InterI" ZF.thy  | 
| 
 
c05fa3ce5439
Improved intersection rule InterI: now truly safe, since the unsafeness is
 
paulson 
parents: 
2716 
diff
changeset
 | 
376  | 
"[| !!x. x: C ==> A: x; EX c. c:C |] ==> A : Inter(C)"  | 
| 4091 | 377  | 
(fn prems=> [ (simp_tac (simpset() addsimps [Inter_iff]) 1),  | 
378  | 
blast_tac (claset() addIs prems) 1 ]);  | 
|
| 2469 | 379  | 
|
380  | 
(*A "destruct" rule -- every B in C contains A as an element, but  | 
|
381  | 
A:B can hold when B:C does not! This rule is analogous to "spec". *)  | 
|
382  | 
qed_goalw "InterD" ZF.thy [Inter_def]  | 
|
383  | 
"!!C. [| A : Inter(C); B : C |] ==> A : B"  | 
|
| 2877 | 384  | 
(fn _=> [ Blast_tac 1 ]);  | 
| 2469 | 385  | 
|
386  | 
(*"Classical" elimination rule -- does not require exhibiting B:C *)  | 
|
387  | 
qed_goalw "InterE" ZF.thy [Inter_def]  | 
|
| 2716 | 388  | 
"[| A : Inter(C); B~:C ==> R; A:B ==> R |] ==> R"  | 
| 2469 | 389  | 
(fn major::prems=>  | 
390  | 
[ (rtac (major RS CollectD2 RS ballE) 1),  | 
|
391  | 
(REPEAT (eresolve_tac prems 1)) ]);  | 
|
392  | 
||
393  | 
AddSIs [InterI];  | 
|
| 2716 | 394  | 
AddEs [InterD, InterE];  | 
| 0 | 395  | 
|
396  | 
(*** Rules for Intersections of families ***)  | 
|
397  | 
(* INT x:A. B(x) abbreviates Inter({B(x). x:A}) *)
 | 
|
398  | 
||
| 2469 | 399  | 
qed_goalw "INT_iff" ZF.thy [Inter_def]  | 
| 485 | 400  | 
"b : (INT x:A. B(x)) <-> (ALL x:A. b : B(x)) & (EX x. x:A)"  | 
| 2469 | 401  | 
(fn _=> [ Simp_tac 1, Best_tac 1 ]);  | 
| 485 | 402  | 
|
| 775 | 403  | 
qed_goal "INT_I" ZF.thy  | 
| 0 | 404  | 
"[| !!x. x: A ==> b: B(x); a: A |] ==> b: (INT x:A. B(x))"  | 
| 4091 | 405  | 
(fn prems=> [ blast_tac (claset() addIs prems) 1 ]);  | 
| 0 | 406  | 
|
| 775 | 407  | 
qed_goal "INT_E" ZF.thy  | 
| 0 | 408  | 
"[| b : (INT x:A. B(x)); a: A |] ==> b : B(a)"  | 
409  | 
(fn [major,minor]=>  | 
|
410  | 
[ (rtac (major RS InterD) 1),  | 
|
411  | 
(rtac (minor RS RepFunI) 1) ]);  | 
|
412  | 
||
| 775 | 413  | 
qed_goal "INT_cong" ZF.thy  | 
| 3840 | 414  | 
"[| A=B; !!x. x:B ==> C(x)=D(x) |] ==> (INT x:A. C(x)) = (INT x:B. D(x))"  | 
| 4091 | 415  | 
(fn prems=> [ (simp_tac (simpset() addsimps prems) 1) ]);  | 
| 2469 | 416  | 
|
417  | 
(*No "Addcongs [INT_cong]" because INT is a combination of constants*)  | 
|
| 435 | 418  | 
|
| 0 | 419  | 
|
420  | 
(*** Rules for Powersets ***)  | 
|
421  | 
||
| 775 | 422  | 
qed_goal "PowI" ZF.thy "A <= B ==> A : Pow(B)"  | 
| 485 | 423  | 
(fn [prem]=> [ (rtac (prem RS (Pow_iff RS iffD2)) 1) ]);  | 
| 0 | 424  | 
|
| 775 | 425  | 
qed_goal "PowD" ZF.thy "A : Pow(B) ==> A<=B"  | 
| 485 | 426  | 
(fn [major]=> [ (rtac (major RS (Pow_iff RS iffD1)) 1) ]);  | 
| 0 | 427  | 
|
| 2469 | 428  | 
AddSIs [PowI];  | 
429  | 
AddSDs [PowD];  | 
|
430  | 
||
| 0 | 431  | 
|
432  | 
(*** Rules for the empty set ***)  | 
|
433  | 
||
434  | 
(*The set {x:0.False} is empty; by foundation it equals 0 
 | 
|
435  | 
See Suppes, page 21.*)  | 
|
| 2469 | 436  | 
qed_goal "not_mem_empty" ZF.thy "a ~: 0"  | 
437  | 
(fn _=>  | 
|
438  | 
[ (cut_facts_tac [foundation] 1),  | 
|
| 4091 | 439  | 
(best_tac (claset() addDs [equalityD2]) 1) ]);  | 
| 2469 | 440  | 
|
441  | 
bind_thm ("emptyE", not_mem_empty RS notE);
 | 
|
442  | 
||
443  | 
Addsimps [not_mem_empty];  | 
|
444  | 
AddSEs [emptyE];  | 
|
| 0 | 445  | 
|
| 775 | 446  | 
qed_goal "empty_subsetI" ZF.thy "0 <= A"  | 
| 2877 | 447  | 
(fn _=> [ Blast_tac 1 ]);  | 
| 2469 | 448  | 
|
449  | 
Addsimps [empty_subsetI];  | 
|
| 0 | 450  | 
|
| 775 | 451  | 
qed_goal "equals0I" ZF.thy "[| !!y. y:A ==> False |] ==> A=0"  | 
| 4091 | 452  | 
(fn prems=> [ blast_tac (claset() addDs prems) 1 ]);  | 
| 0 | 453  | 
|
| 5467 | 454  | 
qed_goal "equals0D" ZF.thy "!!P. A=0 ==> a ~: A"  | 
455  | 
(fn _=> [ Blast_tac 1 ]);  | 
|
| 0 | 456  | 
|
| 5467 | 457  | 
AddDs [equals0D, sym RS equals0D];  | 
| 
5265
 
9d1d4c43c76d
Disjointness reasoning by  AddEs [equals0E, sym RS equals0E]
 
paulson 
parents: 
5242 
diff
changeset
 | 
458  | 
|
| 
825
 
76d9575950f2
Added Krzysztof's theorems subst_elem, not_emptyI, not_emptyE
 
lcp 
parents: 
775 
diff
changeset
 | 
459  | 
qed_goal "not_emptyI" ZF.thy "!!A a. a:A ==> A ~= 0"  | 
| 2877 | 460  | 
(fn _=> [ Blast_tac 1 ]);  | 
| 
825
 
76d9575950f2
Added Krzysztof's theorems subst_elem, not_emptyI, not_emptyE
 
lcp 
parents: 
775 
diff
changeset
 | 
461  | 
|
| 
868
 
452f1e6ae3bc
Deleted semicolon at the end of the qed_goal line, which was preventing
 
lcp 
parents: 
854 
diff
changeset
 | 
462  | 
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
 | 
463  | 
(fn [major,minor]=>  | 
| 
 
76d9575950f2
Added Krzysztof's theorems subst_elem, not_emptyI, not_emptyE
 
lcp 
parents: 
775 
diff
changeset
 | 
464  | 
[ rtac ([major, equals0I] MRS swap) 1,  | 
| 
 
76d9575950f2
Added Krzysztof's theorems subst_elem, not_emptyI, not_emptyE
 
lcp 
parents: 
775 
diff
changeset
 | 
465  | 
swap_res_tac [minor] 1,  | 
| 
 
76d9575950f2
Added Krzysztof's theorems subst_elem, not_emptyI, not_emptyE
 
lcp 
parents: 
775 
diff
changeset
 | 
466  | 
assume_tac 1 ]);  | 
| 
 
76d9575950f2
Added Krzysztof's theorems subst_elem, not_emptyI, not_emptyE
 
lcp 
parents: 
775 
diff
changeset
 | 
467  | 
|
| 0 | 468  | 
|
| 748 | 469  | 
(*** Cantor's Theorem: There is no surjection from a set to its powerset. ***)  | 
470  | 
||
471  | 
val cantor_cs = FOL_cs (*precisely the rules needed for the proof*)  | 
|
472  | 
addSIs [ballI, CollectI, PowI, subsetI] addIs [bexI]  | 
|
473  | 
addSEs [CollectE, equalityCE];  | 
|
474  | 
||
475  | 
(*The search is undirected; similar proof attempts may fail.  | 
|
476  | 
b represents ANY map, such as (lam x:A.b(x)): A->Pow(A). *)  | 
|
| 775 | 477  | 
qed_goal "cantor" ZF.thy "EX S: Pow(A). ALL x:A. b(x) ~= S"  | 
| 2877 | 478  | 
(fn _ => [best_tac cantor_cs 1]);  | 
| 748 | 479  | 
|
| 516 | 480  | 
(*Lemma for the inductive definition in Zorn.thy*)  | 
| 775 | 481  | 
qed_goal "Union_in_Pow" ZF.thy  | 
| 516 | 482  | 
"!!Y. Y : Pow(Pow(A)) ==> Union(Y) : Pow(A)"  | 
| 2877 | 483  | 
(fn _ => [Blast_tac 1]);  | 
| 
1902
 
e349b91cf197
Added function for storing default claset in theory.
 
berghofe 
parents: 
1889 
diff
changeset
 | 
484  | 
|
| 6111 | 485  | 
|
486  | 
local  | 
|
487  | 
val (bspecT, bspec') = make_new_spec bspec  | 
|
488  | 
in  | 
|
489  | 
||
490  | 
fun RSbspec th =  | 
|
491  | 
(case concl_of th of  | 
|
492  | 
     _ $ (Const("Ball",_) $ _ $ Abs(a,_,_)) =>
 | 
|
493  | 
let val ca = cterm_of (#sign(rep_thm th)) (Var((a,0),bspecT))  | 
|
494  | 
in th RS forall_elim ca bspec' end  | 
|
495  | 
  | _ => raise THM("RSbspec",0,[th]));
 | 
|
496  | 
||
497  | 
val normalize_thm_ZF = normalize_thm [RSspec,RSbspec,RSmp];  | 
|
498  | 
||
499  | 
fun qed_spec_mp name =  | 
|
500  | 
let val thm = normalize_thm_ZF (result())  | 
|
501  | 
in bind_thm(name, thm) end;  | 
|
502  | 
||
503  | 
fun qed_goal_spec_mp name thy s p =  | 
|
504  | 
bind_thm (name, normalize_thm_ZF (prove_goal thy s p));  | 
|
505  | 
||
506  | 
fun qed_goalw_spec_mp name thy defs s p =  | 
|
507  | 
bind_thm (name, normalize_thm_ZF (prove_goalw thy defs s p));  | 
|
508  | 
||
509  | 
end;  | 
|
510  | 
||
511  | 
||
512  | 
(* attributes *)  | 
|
513  | 
||
514  | 
local  | 
|
515  | 
||
516  | 
fun gen_rulify x =  | 
|
517  | 
Attrib.no_args (Drule.rule_attribute (K (normalize_thm_ZF))) x;  | 
|
518  | 
||
519  | 
in  | 
|
520  | 
||
521  | 
val attrib_setup =  | 
|
522  | 
[Attrib.add_attributes  | 
|
523  | 
  [("rulify", (gen_rulify, gen_rulify), 
 | 
|
524  | 
"put theorem into standard rule form")]];  | 
|
525  | 
||
526  | 
end;  |