author | gagern |
Wed, 20 Apr 2005 14:18:33 +0200 | |
changeset 15778 | 98af3693f6b3 |
parent 15661 | 9ef583b08647 |
permissions | -rw-r--r-- |
7355 | 1 |
(* Title: FOL/IFOL_lemmas.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1992 University of Cambridge |
|
5 |
||
6 |
Tactics and lemmas for theory IFOL (intuitionistic first-order logic). |
|
7 |
*) |
|
8 |
||
9 |
(* ML bindings *) |
|
10 |
||
11 |
val refl = thm "refl"; |
|
12 |
val subst = thm "subst"; |
|
13 |
val conjI = thm "conjI"; |
|
14 |
val conjunct1 = thm "conjunct1"; |
|
15 |
val conjunct2 = thm "conjunct2"; |
|
16 |
val disjI1 = thm "disjI1"; |
|
17 |
val disjI2 = thm "disjI2"; |
|
18 |
val disjE = thm "disjE"; |
|
19 |
val impI = thm "impI"; |
|
20 |
val mp = thm "mp"; |
|
21 |
val FalseE = thm "FalseE"; |
|
22 |
val True_def = thm "True_def"; |
|
23 |
val not_def = thm "not_def"; |
|
24 |
val iff_def = thm "iff_def"; |
|
25 |
val ex1_def = thm "ex1_def"; |
|
26 |
val allI = thm "allI"; |
|
27 |
val spec = thm "spec"; |
|
28 |
val exI = thm "exI"; |
|
29 |
val exE = thm "exE"; |
|
30 |
val eq_reflection = thm "eq_reflection"; |
|
31 |
val iff_reflection = thm "iff_reflection"; |
|
32 |
||
14085 | 33 |
structure IFOL = |
34 |
struct |
|
35 |
val thy = the_context (); |
|
36 |
val refl = refl; |
|
37 |
val subst = subst; |
|
38 |
val conjI = conjI; |
|
39 |
val conjunct1 = conjunct1; |
|
40 |
val conjunct2 = conjunct2; |
|
41 |
val disjI1 = disjI1; |
|
42 |
val disjI2 = disjI2; |
|
43 |
val disjE = disjE; |
|
44 |
val impI = impI; |
|
45 |
val mp = mp; |
|
46 |
val FalseE = FalseE; |
|
47 |
val True_def = True_def; |
|
48 |
val not_def = not_def; |
|
49 |
val iff_def = iff_def; |
|
50 |
val ex1_def = ex1_def; |
|
51 |
val allI = allI; |
|
52 |
val spec = spec; |
|
53 |
val exI = exI; |
|
54 |
val exE = exE; |
|
55 |
val eq_reflection = eq_reflection; |
|
56 |
val iff_reflection = iff_reflection; |
|
57 |
end; |
|
7355 | 58 |
|
59 |
||
9264 | 60 |
Goalw [True_def] "True"; |
61 |
by (REPEAT (ares_tac [impI] 1)) ; |
|
62 |
qed "TrueI"; |
|
7355 | 63 |
|
64 |
(*** Sequent-style elimination rules for & --> and ALL ***) |
|
65 |
||
9264 | 66 |
val major::prems = Goal |
67 |
"[| P&Q; [| P; Q |] ==> R |] ==> R"; |
|
68 |
by (resolve_tac prems 1); |
|
69 |
by (rtac (major RS conjunct1) 1); |
|
70 |
by (rtac (major RS conjunct2) 1); |
|
71 |
qed "conjE"; |
|
7355 | 72 |
|
9264 | 73 |
val major::prems = Goal |
74 |
"[| P-->Q; P; Q ==> R |] ==> R"; |
|
75 |
by (resolve_tac prems 1); |
|
76 |
by (rtac (major RS mp) 1); |
|
77 |
by (resolve_tac prems 1); |
|
78 |
qed "impE"; |
|
7355 | 79 |
|
9264 | 80 |
val major::prems = Goal |
81 |
"[| ALL x. P(x); P(x) ==> R |] ==> R"; |
|
82 |
by (resolve_tac prems 1); |
|
83 |
by (rtac (major RS spec) 1); |
|
84 |
qed "allE"; |
|
7355 | 85 |
|
86 |
(*Duplicates the quantifier; for use with eresolve_tac*) |
|
9264 | 87 |
val major::prems = Goal |
7355 | 88 |
"[| ALL x. P(x); [| P(x); ALL x. P(x) |] ==> R \ |
9264 | 89 |
\ |] ==> R"; |
90 |
by (resolve_tac prems 1); |
|
91 |
by (rtac (major RS spec) 1); |
|
92 |
by (rtac major 1); |
|
93 |
qed "all_dupE"; |
|
7355 | 94 |
|
95 |
||
96 |
(*** Negation rules, which translate between ~P and P-->False ***) |
|
97 |
||
9264 | 98 |
val prems = Goalw [not_def] "(P ==> False) ==> ~P"; |
99 |
by (REPEAT (ares_tac (prems@[impI]) 1)) ; |
|
100 |
qed "notI"; |
|
7355 | 101 |
|
9264 | 102 |
Goalw [not_def] "[| ~P; P |] ==> R"; |
103 |
by (etac (mp RS FalseE) 1); |
|
104 |
by (assume_tac 1); |
|
105 |
qed "notE"; |
|
7355 | 106 |
|
9264 | 107 |
Goal "[| P; ~P |] ==> R"; |
108 |
by (etac notE 1); |
|
109 |
by (assume_tac 1); |
|
110 |
qed "rev_notE"; |
|
7355 | 111 |
|
112 |
(*This is useful with the special implication rules for each kind of P. *) |
|
9264 | 113 |
val prems = Goal |
114 |
"[| ~P; (P-->False) ==> Q |] ==> Q"; |
|
115 |
by (REPEAT (ares_tac (prems@[impI,notE]) 1)) ; |
|
116 |
qed "not_to_imp"; |
|
7355 | 117 |
|
118 |
(* For substitution into an assumption P, reduce Q to P-->Q, substitute into |
|
119 |
this implication, then apply impI to move P back into the assumptions. |
|
120 |
To specify P use something like |
|
121 |
eres_inst_tac [ ("P","ALL y. ?S(x,y)") ] rev_mp 1 *) |
|
9264 | 122 |
Goal "[| P; P --> Q |] ==> Q"; |
123 |
by (etac mp 1); |
|
124 |
by (assume_tac 1); |
|
125 |
qed "rev_mp"; |
|
7355 | 126 |
|
127 |
(*Contrapositive of an inference rule*) |
|
9264 | 128 |
val [major,minor]= Goal "[| ~Q; P==>Q |] ==> ~P"; |
129 |
by (rtac (major RS notE RS notI) 1); |
|
130 |
by (etac minor 1) ; |
|
131 |
qed "contrapos"; |
|
7355 | 132 |
|
133 |
||
134 |
(*** Modus Ponens Tactics ***) |
|
135 |
||
136 |
(*Finds P-->Q and P in the assumptions, replaces implication by Q *) |
|
137 |
fun mp_tac i = eresolve_tac [notE,impE] i THEN assume_tac i; |
|
138 |
||
139 |
(*Like mp_tac but instantiates no variables*) |
|
140 |
fun eq_mp_tac i = eresolve_tac [notE,impE] i THEN eq_assume_tac i; |
|
141 |
||
142 |
||
143 |
(*** If-and-only-if ***) |
|
144 |
||
9264 | 145 |
val prems = Goalw [iff_def] |
146 |
"[| P ==> Q; Q ==> P |] ==> P<->Q"; |
|
147 |
by (REPEAT (ares_tac (prems@[conjI, impI]) 1)) ; |
|
148 |
qed "iffI"; |
|
7355 | 149 |
|
150 |
||
151 |
(*Observe use of rewrite_rule to unfold "<->" in meta-assumptions (prems) *) |
|
9264 | 152 |
val prems = Goalw [iff_def] |
153 |
"[| P <-> Q; [| P-->Q; Q-->P |] ==> R |] ==> R"; |
|
154 |
by (rtac conjE 1); |
|
155 |
by (REPEAT (ares_tac prems 1)) ; |
|
156 |
qed "iffE"; |
|
7355 | 157 |
|
158 |
(* Destruct rules for <-> similar to Modus Ponens *) |
|
159 |
||
9264 | 160 |
Goalw [iff_def] "[| P <-> Q; P |] ==> Q"; |
161 |
by (etac (conjunct1 RS mp) 1); |
|
162 |
by (assume_tac 1); |
|
163 |
qed "iffD1"; |
|
7355 | 164 |
|
9264 | 165 |
val prems = Goalw [iff_def] "[| P <-> Q; Q |] ==> P"; |
166 |
by (etac (conjunct2 RS mp) 1); |
|
167 |
by (assume_tac 1); |
|
168 |
qed "iffD2"; |
|
7355 | 169 |
|
9264 | 170 |
Goal "[| P; P <-> Q |] ==> Q"; |
171 |
by (etac iffD1 1); |
|
172 |
by (assume_tac 1); |
|
173 |
qed "rev_iffD1"; |
|
7355 | 174 |
|
9264 | 175 |
Goal "[| Q; P <-> Q |] ==> P"; |
176 |
by (etac iffD2 1); |
|
177 |
by (assume_tac 1); |
|
178 |
qed "rev_iffD2"; |
|
179 |
||
180 |
Goal "P <-> P"; |
|
181 |
by (REPEAT (ares_tac [iffI] 1)) ; |
|
182 |
qed "iff_refl"; |
|
7355 | 183 |
|
9264 | 184 |
Goal "Q <-> P ==> P <-> Q"; |
185 |
by (etac iffE 1); |
|
186 |
by (rtac iffI 1); |
|
187 |
by (REPEAT (eresolve_tac [asm_rl,mp] 1)) ; |
|
188 |
qed "iff_sym"; |
|
7355 | 189 |
|
9264 | 190 |
Goal "[| P <-> Q; Q<-> R |] ==> P <-> R"; |
191 |
by (rtac iffI 1); |
|
192 |
by (REPEAT (eresolve_tac [asm_rl,iffE] 1 ORELSE mp_tac 1)) ; |
|
193 |
qed "iff_trans"; |
|
7355 | 194 |
|
195 |
||
196 |
(*** Unique existence. NOTE THAT the following 2 quantifications |
|
197 |
EX!x such that [EX!y such that P(x,y)] (sequential) |
|
198 |
EX!x,y such that P(x,y) (simultaneous) |
|
199 |
do NOT mean the same thing. The parser treats EX!x y.P(x,y) as sequential. |
|
200 |
***) |
|
201 |
||
9264 | 202 |
val prems = Goalw [ex1_def] |
203 |
"[| P(a); !!x. P(x) ==> x=a |] ==> EX! x. P(x)"; |
|
204 |
by (REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)) ; |
|
205 |
qed "ex1I"; |
|
7355 | 206 |
|
15661
9ef583b08647
reverted renaming of Some/None in comments and strings;
wenzelm
parents:
15570
diff
changeset
|
207 |
(*Sometimes easier to use: the premises have no shared variables. Safe!*) |
9264 | 208 |
val [ex,eq] = Goal |
209 |
"[| EX x. P(x); !!x y. [| P(x); P(y) |] ==> x=y |] ==> EX! x. P(x)"; |
|
210 |
by (rtac (ex RS exE) 1); |
|
211 |
by (REPEAT (ares_tac [ex1I,eq] 1)) ; |
|
212 |
qed "ex_ex1I"; |
|
7355 | 213 |
|
9264 | 214 |
val prems = Goalw [ex1_def] |
215 |
"[| EX! x. P(x); !!x. [| P(x); ALL y. P(y) --> y=x |] ==> R |] ==> R"; |
|
216 |
by (cut_facts_tac prems 1); |
|
217 |
by (REPEAT (eresolve_tac [exE,conjE] 1 ORELSE ares_tac prems 1)) ; |
|
218 |
qed "ex1E"; |
|
7355 | 219 |
|
220 |
||
221 |
(*** <-> congruence rules for simplification ***) |
|
222 |
||
223 |
(*Use iffE on a premise. For conj_cong, imp_cong, all_cong, ex_cong*) |
|
224 |
fun iff_tac prems i = |
|
225 |
resolve_tac (prems RL [iffE]) i THEN |
|
226 |
REPEAT1 (eresolve_tac [asm_rl,mp] i); |
|
227 |
||
9264 | 228 |
val prems = Goal |
229 |
"[| P <-> P'; P' ==> Q <-> Q' |] ==> (P&Q) <-> (P'&Q')"; |
|
230 |
by (cut_facts_tac prems 1); |
|
231 |
by (REPEAT (ares_tac [iffI,conjI] 1 |
|
232 |
ORELSE eresolve_tac [iffE,conjE,mp] 1 |
|
233 |
ORELSE iff_tac prems 1)) ; |
|
234 |
qed "conj_cong"; |
|
7355 | 235 |
|
236 |
(*Reversed congruence rule! Used in ZF/Order*) |
|
9264 | 237 |
val prems = Goal |
238 |
"[| P <-> P'; P' ==> Q <-> Q' |] ==> (Q&P) <-> (Q'&P')"; |
|
239 |
by (cut_facts_tac prems 1); |
|
240 |
by (REPEAT (ares_tac [iffI,conjI] 1 |
|
241 |
ORELSE eresolve_tac [iffE,conjE,mp] 1 ORELSE iff_tac prems 1)) ; |
|
242 |
qed "conj_cong2"; |
|
7355 | 243 |
|
9264 | 244 |
val prems = Goal |
245 |
"[| P <-> P'; Q <-> Q' |] ==> (P|Q) <-> (P'|Q')"; |
|
246 |
by (cut_facts_tac prems 1); |
|
247 |
by (REPEAT (eresolve_tac [iffE,disjE,disjI1,disjI2] 1 |
|
248 |
ORELSE ares_tac [iffI] 1 ORELSE mp_tac 1)) ; |
|
249 |
qed "disj_cong"; |
|
7355 | 250 |
|
9264 | 251 |
val prems = Goal |
252 |
"[| P <-> P'; P' ==> Q <-> Q' |] ==> (P-->Q) <-> (P'-->Q')"; |
|
253 |
by (cut_facts_tac prems 1); |
|
254 |
by (REPEAT (ares_tac [iffI,impI] 1 |
|
255 |
ORELSE etac iffE 1 ORELSE mp_tac 1 ORELSE iff_tac prems 1)) ; |
|
256 |
qed "imp_cong"; |
|
7355 | 257 |
|
9264 | 258 |
val prems = Goal |
259 |
"[| P <-> P'; Q <-> Q' |] ==> (P<->Q) <-> (P'<->Q')"; |
|
260 |
by (cut_facts_tac prems 1); |
|
261 |
by (REPEAT (etac iffE 1 ORELSE ares_tac [iffI] 1 ORELSE mp_tac 1)) ; |
|
262 |
qed "iff_cong"; |
|
7355 | 263 |
|
9264 | 264 |
val prems = Goal |
265 |
"P <-> P' ==> ~P <-> ~P'"; |
|
266 |
by (cut_facts_tac prems 1); |
|
267 |
by (REPEAT (ares_tac [iffI,notI] 1 |
|
268 |
ORELSE mp_tac 1 ORELSE eresolve_tac [iffE,notE] 1)) ; |
|
269 |
qed "not_cong"; |
|
7355 | 270 |
|
9264 | 271 |
val prems = Goal |
272 |
"(!!x. P(x) <-> Q(x)) ==> (ALL x. P(x)) <-> (ALL x. Q(x))"; |
|
273 |
by (REPEAT (ares_tac [iffI,allI] 1 |
|
274 |
ORELSE mp_tac 1 ORELSE etac allE 1 ORELSE iff_tac prems 1)) ; |
|
275 |
qed "all_cong"; |
|
7355 | 276 |
|
9264 | 277 |
val prems = Goal |
278 |
"(!!x. P(x) <-> Q(x)) ==> (EX x. P(x)) <-> (EX x. Q(x))"; |
|
279 |
by (REPEAT (etac exE 1 ORELSE ares_tac [iffI,exI] 1 |
|
280 |
ORELSE mp_tac 1 ORELSE iff_tac prems 1)) ; |
|
281 |
qed "ex_cong"; |
|
7355 | 282 |
|
9264 | 283 |
val prems = Goal |
284 |
"(!!x. P(x) <-> Q(x)) ==> (EX! x. P(x)) <-> (EX! x. Q(x))"; |
|
285 |
by (REPEAT (eresolve_tac [ex1E, spec RS mp] 1 |
|
286 |
ORELSE ares_tac [iffI,ex1I] 1 ORELSE mp_tac 1 |
|
287 |
ORELSE iff_tac prems 1)) ; |
|
288 |
qed "ex1_cong"; |
|
7355 | 289 |
|
290 |
(*** Equality rules ***) |
|
291 |
||
9264 | 292 |
Goal "a=b ==> b=a"; |
293 |
by (etac subst 1); |
|
294 |
by (rtac refl 1) ; |
|
295 |
qed "sym"; |
|
7355 | 296 |
|
9264 | 297 |
Goal "[| a=b; b=c |] ==> a=c"; |
298 |
by (etac subst 1 THEN assume_tac 1) ; |
|
299 |
qed "trans"; |
|
7355 | 300 |
|
301 |
(** ~ b=a ==> ~ a=b **) |
|
7422 | 302 |
bind_thm ("not_sym", hd (compose(sym,2,contrapos))); |
7355 | 303 |
|
304 |
||
305 |
(* Two theorms for rewriting only one instance of a definition: |
|
306 |
the first for definitions of formulae and the second for terms *) |
|
307 |
||
9264 | 308 |
val prems = goal (the_context()) "(A == B) ==> A <-> B"; |
7355 | 309 |
by (rewrite_goals_tac prems); |
310 |
by (rtac iff_refl 1); |
|
311 |
qed "def_imp_iff"; |
|
312 |
||
9264 | 313 |
val prems = goal (the_context()) "(A == B) ==> A = B"; |
7355 | 314 |
by (rewrite_goals_tac prems); |
315 |
by (rtac refl 1); |
|
316 |
qed "meta_eq_to_obj_eq"; |
|
317 |
||
9527 | 318 |
(*substitution*) |
7355 | 319 |
bind_thm ("ssubst", sym RS subst); |
320 |
||
321 |
(*A special case of ex1E that would otherwise need quantifier expansion*) |
|
9264 | 322 |
val prems = Goal |
323 |
"[| EX! x. P(x); P(a); P(b) |] ==> a=b"; |
|
324 |
by (cut_facts_tac prems 1); |
|
325 |
by (etac ex1E 1); |
|
326 |
by (rtac trans 1); |
|
327 |
by (rtac sym 2); |
|
328 |
by (REPEAT (eresolve_tac [asm_rl, spec RS mp] 1)) ; |
|
329 |
qed "ex1_equalsE"; |
|
7355 | 330 |
|
331 |
(** Polymorphic congruence rules **) |
|
332 |
||
9264 | 333 |
Goal "[| a=b |] ==> t(a)=t(b)"; |
334 |
by (etac ssubst 1); |
|
335 |
by (rtac refl 1) ; |
|
336 |
qed "subst_context"; |
|
7355 | 337 |
|
9264 | 338 |
Goal "[| a=b; c=d |] ==> t(a,c)=t(b,d)"; |
339 |
by (REPEAT (etac ssubst 1)); |
|
340 |
by (rtac refl 1) ; |
|
341 |
qed "subst_context2"; |
|
7355 | 342 |
|
9264 | 343 |
Goal "[| a=b; c=d; e=f |] ==> t(a,c,e)=t(b,d,f)"; |
344 |
by (REPEAT (etac ssubst 1)); |
|
345 |
by (rtac refl 1) ; |
|
346 |
qed "subst_context3"; |
|
7355 | 347 |
|
348 |
(*Useful with eresolve_tac for proving equalties from known equalities. |
|
349 |
a = b |
|
350 |
| | |
|
351 |
c = d *) |
|
9264 | 352 |
Goal "[| a=b; a=c; b=d |] ==> c=d"; |
353 |
by (rtac trans 1); |
|
354 |
by (rtac trans 1); |
|
355 |
by (rtac sym 1); |
|
356 |
by (REPEAT (assume_tac 1)); |
|
357 |
qed "box_equals"; |
|
7355 | 358 |
|
359 |
(*Dual of box_equals: for proving equalities backwards*) |
|
9264 | 360 |
Goal "[| a=c; b=d; c=d |] ==> a=b"; |
361 |
by (rtac trans 1); |
|
362 |
by (rtac trans 1); |
|
363 |
by (REPEAT (assume_tac 1)); |
|
364 |
by (etac sym 1); |
|
365 |
qed "simp_equals"; |
|
7355 | 366 |
|
367 |
(** Congruence rules for predicate letters **) |
|
368 |
||
9264 | 369 |
Goal "a=a' ==> P(a) <-> P(a')"; |
370 |
by (rtac iffI 1); |
|
371 |
by (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ; |
|
372 |
qed "pred1_cong"; |
|
7355 | 373 |
|
9264 | 374 |
Goal "[| a=a'; b=b' |] ==> P(a,b) <-> P(a',b')"; |
375 |
by (rtac iffI 1); |
|
376 |
by (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ; |
|
377 |
qed "pred2_cong"; |
|
7355 | 378 |
|
9264 | 379 |
Goal "[| a=a'; b=b'; c=c' |] ==> P(a,b,c) <-> P(a',b',c')"; |
380 |
by (rtac iffI 1); |
|
381 |
by (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ; |
|
382 |
qed "pred3_cong"; |
|
7355 | 383 |
|
384 |
(*special cases for free variables P, Q, R, S -- up to 3 arguments*) |
|
385 |
||
386 |
val pred_congs = |
|
15570 | 387 |
List.concat (map (fn c => |
7355 | 388 |
map (fn th => read_instantiate [("P",c)] th) |
389 |
[pred1_cong,pred2_cong,pred3_cong]) |
|
390 |
(explode"PQRS")); |
|
391 |
||
392 |
(*special case for the equality predicate!*) |
|
7422 | 393 |
bind_thm ("eq_cong", read_instantiate [("P","op =")] pred2_cong); |
7355 | 394 |
|
395 |
||
396 |
(*** Simplifications of assumed implications. |
|
397 |
Roy Dyckhoff has proved that conj_impE, disj_impE, and imp_impE |
|
398 |
used with mp_tac (restricted to atomic formulae) is COMPLETE for |
|
399 |
intuitionistic propositional logic. See |
|
400 |
R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic |
|
401 |
(preprint, University of St Andrews, 1991) ***) |
|
402 |
||
9264 | 403 |
val major::prems= Goal |
404 |
"[| (P&Q)-->S; P-->(Q-->S) ==> R |] ==> R"; |
|
405 |
by (REPEAT (ares_tac ([conjI, impI, major RS mp]@prems) 1)) ; |
|
406 |
qed "conj_impE"; |
|
7355 | 407 |
|
9264 | 408 |
val major::prems= Goal |
409 |
"[| (P|Q)-->S; [| P-->S; Q-->S |] ==> R |] ==> R"; |
|
410 |
by (DEPTH_SOLVE (ares_tac ([disjI1, disjI2, impI, major RS mp]@prems) 1)) ; |
|
411 |
qed "disj_impE"; |
|
7355 | 412 |
|
413 |
(*Simplifies the implication. Classical version is stronger. |
|
414 |
Still UNSAFE since Q must be provable -- backtracking needed. *) |
|
9264 | 415 |
val major::prems= Goal |
416 |
"[| (P-->Q)-->S; [| P; Q-->S |] ==> Q; S ==> R |] ==> R"; |
|
417 |
by (REPEAT (ares_tac ([impI, major RS mp]@prems) 1)) ; |
|
418 |
qed "imp_impE"; |
|
7355 | 419 |
|
420 |
(*Simplifies the implication. Classical version is stronger. |
|
421 |
Still UNSAFE since ~P must be provable -- backtracking needed. *) |
|
9264 | 422 |
val major::prems= Goal |
423 |
"[| ~P --> S; P ==> False; S ==> R |] ==> R"; |
|
424 |
by (REPEAT (ares_tac ([notI, impI, major RS mp]@prems) 1)) ; |
|
425 |
qed "not_impE"; |
|
7355 | 426 |
|
427 |
(*Simplifies the implication. UNSAFE. *) |
|
9264 | 428 |
val major::prems= Goal |
7355 | 429 |
"[| (P<->Q)-->S; [| P; Q-->S |] ==> Q; [| Q; P-->S |] ==> P; \ |
9264 | 430 |
\ S ==> R |] ==> R"; |
431 |
by (REPEAT (ares_tac ([iffI, impI, major RS mp]@prems) 1)) ; |
|
432 |
qed "iff_impE"; |
|
7355 | 433 |
|
434 |
(*What if (ALL x.~~P(x)) --> ~~(ALL x.P(x)) is an assumption? UNSAFE*) |
|
9264 | 435 |
val major::prems= Goal |
436 |
"[| (ALL x. P(x))-->S; !!x. P(x); S ==> R |] ==> R"; |
|
437 |
by (REPEAT (ares_tac ([allI, impI, major RS mp]@prems) 1)) ; |
|
438 |
qed "all_impE"; |
|
7355 | 439 |
|
440 |
(*Unsafe: (EX x.P(x))-->S is equivalent to ALL x.P(x)-->S. *) |
|
9264 | 441 |
val major::prems= Goal |
442 |
"[| (EX x. P(x))-->S; P(x)-->S ==> R |] ==> R"; |
|
443 |
by (REPEAT (ares_tac ([exI, impI, major RS mp]@prems) 1)) ; |
|
444 |
qed "ex_impE"; |
|
7355 | 445 |
|
446 |
(*** Courtesy of Krzysztof Grabczewski ***) |
|
447 |
||
9264 | 448 |
val major::prems = Goal "[| P|Q; P==>R; Q==>S |] ==> R|S"; |
7355 | 449 |
by (rtac (major RS disjE) 1); |
450 |
by (REPEAT (eresolve_tac (prems RL [disjI1, disjI2]) 1)); |
|
451 |
qed "disj_imp_disj"; |