author | clasohm |
Tue, 13 Dec 1994 11:51:12 +0100 | |
changeset 779 | 4ab9176b45b7 |
parent 12 | f17d542276b6 |
child 793 | 0b5c5f568d74 |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: FOL/ifol.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 ifol.thy (intuitionistic first-order logic) |
|
7 |
*) |
|
8 |
||
9 |
open IFOL; |
|
10 |
||
11 |
||
779 | 12 |
qed_goalw "TrueI" IFOL.thy [True_def] "True" |
0 | 13 |
(fn _ => [ (REPEAT (ares_tac [impI] 1)) ]); |
14 |
||
15 |
(*** Sequent-style elimination rules for & --> and ALL ***) |
|
16 |
||
779 | 17 |
qed_goal "conjE" IFOL.thy |
0 | 18 |
"[| P&Q; [| P; Q |] ==> R |] ==> R" |
19 |
(fn prems=> |
|
20 |
[ (REPEAT (resolve_tac prems 1 |
|
21 |
ORELSE (resolve_tac [conjunct1, conjunct2] 1 THEN |
|
22 |
resolve_tac prems 1))) ]); |
|
23 |
||
779 | 24 |
qed_goal "impE" IFOL.thy |
0 | 25 |
"[| P-->Q; P; Q ==> R |] ==> R" |
26 |
(fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); |
|
27 |
||
779 | 28 |
qed_goal "allE" IFOL.thy |
0 | 29 |
"[| ALL x.P(x); P(x) ==> R |] ==> R" |
30 |
(fn prems=> [ (REPEAT (resolve_tac (prems@[spec]) 1)) ]); |
|
31 |
||
32 |
(*Duplicates the quantifier; for use with eresolve_tac*) |
|
779 | 33 |
qed_goal "all_dupE" IFOL.thy |
0 | 34 |
"[| ALL x.P(x); [| P(x); ALL x.P(x) |] ==> R \ |
35 |
\ |] ==> R" |
|
36 |
(fn prems=> [ (REPEAT (resolve_tac (prems@[spec]) 1)) ]); |
|
37 |
||
38 |
||
39 |
(*** Negation rules, which translate between ~P and P-->False ***) |
|
40 |
||
779 | 41 |
qed_goalw "notI" IFOL.thy [not_def] "(P ==> False) ==> ~P" |
0 | 42 |
(fn prems=> [ (REPEAT (ares_tac (prems@[impI]) 1)) ]); |
43 |
||
779 | 44 |
qed_goalw "notE" IFOL.thy [not_def] "[| ~P; P |] ==> R" |
0 | 45 |
(fn prems=> |
46 |
[ (resolve_tac [mp RS FalseE] 1), |
|
47 |
(REPEAT (resolve_tac prems 1)) ]); |
|
48 |
||
49 |
(*This is useful with the special implication rules for each kind of P. *) |
|
779 | 50 |
qed_goal "not_to_imp" IFOL.thy |
0 | 51 |
"[| ~P; (P-->False) ==> Q |] ==> Q" |
52 |
(fn prems=> [ (REPEAT (ares_tac (prems@[impI,notE]) 1)) ]); |
|
53 |
||
54 |
||
55 |
(* For substitution int an assumption P, reduce Q to P-->Q, substitute into |
|
56 |
this implication, then apply impI to move P back into the assumptions. |
|
57 |
To specify P use something like |
|
58 |
eres_inst_tac [ ("P","ALL y. ?S(x,y)") ] rev_mp 1 *) |
|
779 | 59 |
qed_goal "rev_mp" IFOL.thy "[| P; P --> Q |] ==> Q" |
0 | 60 |
(fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); |
61 |
||
62 |
||
63 |
(*Contrapositive of an inference rule*) |
|
779 | 64 |
qed_goal "contrapos" IFOL.thy "[| ~Q; P==>Q |] ==> ~P" |
0 | 65 |
(fn [major,minor]=> |
66 |
[ (rtac (major RS notE RS notI) 1), |
|
67 |
(etac minor 1) ]); |
|
68 |
||
69 |
||
70 |
(*** Modus Ponens Tactics ***) |
|
71 |
||
72 |
(*Finds P-->Q and P in the assumptions, replaces implication by Q *) |
|
73 |
fun mp_tac i = eresolve_tac [notE,impE] i THEN assume_tac i; |
|
74 |
||
75 |
(*Like mp_tac but instantiates no variables*) |
|
76 |
fun eq_mp_tac i = eresolve_tac [notE,impE] i THEN eq_assume_tac i; |
|
77 |
||
78 |
||
79 |
(*** If-and-only-if ***) |
|
80 |
||
779 | 81 |
qed_goalw "iffI" IFOL.thy [iff_def] |
0 | 82 |
"[| P ==> Q; Q ==> P |] ==> P<->Q" |
83 |
(fn prems=> [ (REPEAT (ares_tac (prems@[conjI, impI]) 1)) ]); |
|
84 |
||
85 |
||
86 |
(*Observe use of rewrite_rule to unfold "<->" in meta-assumptions (prems) *) |
|
779 | 87 |
qed_goalw "iffE" IFOL.thy [iff_def] |
0 | 88 |
"[| P <-> Q; [| P-->Q; Q-->P |] ==> R |] ==> R" |
89 |
(fn prems => [ (resolve_tac [conjE] 1), (REPEAT (ares_tac prems 1)) ]); |
|
90 |
||
91 |
(* Destruct rules for <-> similar to Modus Ponens *) |
|
92 |
||
779 | 93 |
qed_goalw "iffD1" IFOL.thy [iff_def] "[| P <-> Q; P |] ==> Q" |
0 | 94 |
(fn prems => [ (rtac (conjunct1 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); |
95 |
||
779 | 96 |
qed_goalw "iffD2" IFOL.thy [iff_def] "[| P <-> Q; Q |] ==> P" |
0 | 97 |
(fn prems => [ (rtac (conjunct2 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); |
98 |
||
779 | 99 |
qed_goal "iff_refl" IFOL.thy "P <-> P" |
0 | 100 |
(fn _ => [ (REPEAT (ares_tac [iffI] 1)) ]); |
101 |
||
779 | 102 |
qed_goal "iff_sym" IFOL.thy "Q <-> P ==> P <-> Q" |
0 | 103 |
(fn [major] => |
104 |
[ (rtac (major RS iffE) 1), |
|
105 |
(rtac iffI 1), |
|
106 |
(REPEAT (eresolve_tac [asm_rl,mp] 1)) ]); |
|
107 |
||
779 | 108 |
qed_goal "iff_trans" IFOL.thy |
0 | 109 |
"!!P Q R. [| P <-> Q; Q<-> R |] ==> P <-> R" |
110 |
(fn _ => |
|
111 |
[ (rtac iffI 1), |
|
112 |
(REPEAT (eresolve_tac [asm_rl,iffE] 1 ORELSE mp_tac 1)) ]); |
|
113 |
||
114 |
||
115 |
(*** Unique existence. NOTE THAT the following 2 quantifications |
|
116 |
EX!x such that [EX!y such that P(x,y)] (sequential) |
|
117 |
EX!x,y such that P(x,y) (simultaneous) |
|
118 |
do NOT mean the same thing. The parser treats EX!x y.P(x,y) as sequential. |
|
119 |
***) |
|
120 |
||
779 | 121 |
qed_goalw "ex1I" IFOL.thy [ex1_def] |
0 | 122 |
"[| P(a); !!x. P(x) ==> x=a |] ==> EX! x. P(x)" |
123 |
(fn prems => [ (REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)) ]); |
|
124 |
||
12
f17d542276b6
Added ex_ex1I: new introduction rule for the EX! quantifier.
lcp
parents:
0
diff
changeset
|
125 |
(*Sometimes easier to use: the premises have no shared variables*) |
779 | 126 |
qed_goal "ex_ex1I" IFOL.thy |
12
f17d542276b6
Added ex_ex1I: new introduction rule for the EX! quantifier.
lcp
parents:
0
diff
changeset
|
127 |
"[| EX x.P(x); !!x y. [| P(x); P(y) |] ==> x=y |] ==> EX! x. P(x)" |
f17d542276b6
Added ex_ex1I: new introduction rule for the EX! quantifier.
lcp
parents:
0
diff
changeset
|
128 |
(fn [ex,eq] => [ (rtac (ex RS exE) 1), |
f17d542276b6
Added ex_ex1I: new introduction rule for the EX! quantifier.
lcp
parents:
0
diff
changeset
|
129 |
(REPEAT (ares_tac [ex1I,eq] 1)) ]); |
f17d542276b6
Added ex_ex1I: new introduction rule for the EX! quantifier.
lcp
parents:
0
diff
changeset
|
130 |
|
779 | 131 |
qed_goalw "ex1E" IFOL.thy [ex1_def] |
0 | 132 |
"[| EX! x.P(x); !!x. [| P(x); ALL y. P(y) --> y=x |] ==> R |] ==> R" |
133 |
(fn prems => |
|
134 |
[ (cut_facts_tac prems 1), |
|
135 |
(REPEAT (eresolve_tac [exE,conjE] 1 ORELSE ares_tac prems 1)) ]); |
|
136 |
||
137 |
||
138 |
(*** <-> congruence rules for simplification ***) |
|
139 |
||
140 |
(*Use iffE on a premise. For conj_cong, imp_cong, all_cong, ex_cong*) |
|
141 |
fun iff_tac prems i = |
|
142 |
resolve_tac (prems RL [iffE]) i THEN |
|
143 |
REPEAT1 (eresolve_tac [asm_rl,mp] i); |
|
144 |
||
779 | 145 |
qed_goal "conj_cong" IFOL.thy |
0 | 146 |
"[| P <-> P'; P' ==> Q <-> Q' |] ==> (P&Q) <-> (P'&Q')" |
147 |
(fn prems => |
|
148 |
[ (cut_facts_tac prems 1), |
|
149 |
(REPEAT (ares_tac [iffI,conjI] 1 |
|
150 |
ORELSE eresolve_tac [iffE,conjE,mp] 1 |
|
151 |
ORELSE iff_tac prems 1)) ]); |
|
152 |
||
779 | 153 |
qed_goal "disj_cong" IFOL.thy |
0 | 154 |
"[| P <-> P'; Q <-> Q' |] ==> (P|Q) <-> (P'|Q')" |
155 |
(fn prems => |
|
156 |
[ (cut_facts_tac prems 1), |
|
157 |
(REPEAT (eresolve_tac [iffE,disjE,disjI1,disjI2] 1 |
|
158 |
ORELSE ares_tac [iffI] 1 |
|
159 |
ORELSE mp_tac 1)) ]); |
|
160 |
||
779 | 161 |
qed_goal "imp_cong" IFOL.thy |
0 | 162 |
"[| P <-> P'; P' ==> Q <-> Q' |] ==> (P-->Q) <-> (P'-->Q')" |
163 |
(fn prems => |
|
164 |
[ (cut_facts_tac prems 1), |
|
165 |
(REPEAT (ares_tac [iffI,impI] 1 |
|
166 |
ORELSE eresolve_tac [iffE] 1 |
|
167 |
ORELSE mp_tac 1 ORELSE iff_tac prems 1)) ]); |
|
168 |
||
779 | 169 |
qed_goal "iff_cong" IFOL.thy |
0 | 170 |
"[| P <-> P'; Q <-> Q' |] ==> (P<->Q) <-> (P'<->Q')" |
171 |
(fn prems => |
|
172 |
[ (cut_facts_tac prems 1), |
|
173 |
(REPEAT (eresolve_tac [iffE] 1 |
|
174 |
ORELSE ares_tac [iffI] 1 |
|
175 |
ORELSE mp_tac 1)) ]); |
|
176 |
||
779 | 177 |
qed_goal "not_cong" IFOL.thy |
0 | 178 |
"P <-> P' ==> ~P <-> ~P'" |
179 |
(fn prems => |
|
180 |
[ (cut_facts_tac prems 1), |
|
181 |
(REPEAT (ares_tac [iffI,notI] 1 |
|
182 |
ORELSE mp_tac 1 |
|
183 |
ORELSE eresolve_tac [iffE,notE] 1)) ]); |
|
184 |
||
779 | 185 |
qed_goal "all_cong" IFOL.thy |
0 | 186 |
"(!!x.P(x) <-> Q(x)) ==> (ALL x.P(x)) <-> (ALL x.Q(x))" |
187 |
(fn prems => |
|
188 |
[ (REPEAT (ares_tac [iffI,allI] 1 |
|
189 |
ORELSE mp_tac 1 |
|
190 |
ORELSE eresolve_tac [allE] 1 ORELSE iff_tac prems 1)) ]); |
|
191 |
||
779 | 192 |
qed_goal "ex_cong" IFOL.thy |
0 | 193 |
"(!!x.P(x) <-> Q(x)) ==> (EX x.P(x)) <-> (EX x.Q(x))" |
194 |
(fn prems => |
|
195 |
[ (REPEAT (eresolve_tac [exE] 1 ORELSE ares_tac [iffI,exI] 1 |
|
196 |
ORELSE mp_tac 1 |
|
197 |
ORELSE iff_tac prems 1)) ]); |
|
198 |
||
779 | 199 |
qed_goal "ex1_cong" IFOL.thy |
0 | 200 |
"(!!x.P(x) <-> Q(x)) ==> (EX! x.P(x)) <-> (EX! x.Q(x))" |
201 |
(fn prems => |
|
202 |
[ (REPEAT (eresolve_tac [ex1E, spec RS mp] 1 ORELSE ares_tac [iffI,ex1I] 1 |
|
203 |
ORELSE mp_tac 1 |
|
204 |
ORELSE iff_tac prems 1)) ]); |
|
205 |
||
206 |
(*** Equality rules ***) |
|
207 |
||
779 | 208 |
qed_goal "sym" IFOL.thy "a=b ==> b=a" |
0 | 209 |
(fn [major] => [ (rtac (major RS subst) 1), (rtac refl 1) ]); |
210 |
||
779 | 211 |
qed_goal "trans" IFOL.thy "[| a=b; b=c |] ==> a=c" |
0 | 212 |
(fn [prem1,prem2] => [ (rtac (prem2 RS subst) 1), (rtac prem1 1) ]); |
213 |
||
214 |
(** ~ b=a ==> ~ a=b **) |
|
215 |
val [not_sym] = compose(sym,2,contrapos); |
|
216 |
||
217 |
(*calling "standard" reduces maxidx to 0*) |
|
779 | 218 |
bind_thm ("ssubst", (sym RS subst)); |
0 | 219 |
|
220 |
(*A special case of ex1E that would otherwise need quantifier expansion*) |
|
779 | 221 |
qed_goal "ex1_equalsE" IFOL.thy |
0 | 222 |
"[| EX! x.P(x); P(a); P(b) |] ==> a=b" |
223 |
(fn prems => |
|
224 |
[ (cut_facts_tac prems 1), |
|
225 |
(etac ex1E 1), |
|
226 |
(rtac trans 1), |
|
227 |
(rtac sym 2), |
|
228 |
(REPEAT (eresolve_tac [asm_rl, spec RS mp] 1)) ]); |
|
229 |
||
230 |
(** Polymorphic congruence rules **) |
|
231 |
||
779 | 232 |
qed_goal "subst_context" IFOL.thy |
0 | 233 |
"[| a=b |] ==> t(a)=t(b)" |
234 |
(fn prems=> |
|
235 |
[ (resolve_tac (prems RL [ssubst]) 1), |
|
236 |
(resolve_tac [refl] 1) ]); |
|
237 |
||
779 | 238 |
qed_goal "subst_context2" IFOL.thy |
0 | 239 |
"[| a=b; c=d |] ==> t(a,c)=t(b,d)" |
240 |
(fn prems=> |
|
241 |
[ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); |
|
242 |
||
779 | 243 |
qed_goal "subst_context3" IFOL.thy |
0 | 244 |
"[| a=b; c=d; e=f |] ==> t(a,c,e)=t(b,d,f)" |
245 |
(fn prems=> |
|
246 |
[ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); |
|
247 |
||
248 |
(*Useful with eresolve_tac for proving equalties from known equalities. |
|
249 |
a = b |
|
250 |
| | |
|
251 |
c = d *) |
|
779 | 252 |
qed_goal "box_equals" IFOL.thy |
0 | 253 |
"[| a=b; a=c; b=d |] ==> c=d" |
254 |
(fn prems=> |
|
255 |
[ (resolve_tac [trans] 1), |
|
256 |
(resolve_tac [trans] 1), |
|
257 |
(resolve_tac [sym] 1), |
|
258 |
(REPEAT (resolve_tac prems 1)) ]); |
|
259 |
||
260 |
(*Dual of box_equals: for proving equalities backwards*) |
|
779 | 261 |
qed_goal "simp_equals" IFOL.thy |
0 | 262 |
"[| a=c; b=d; c=d |] ==> a=b" |
263 |
(fn prems=> |
|
264 |
[ (resolve_tac [trans] 1), |
|
265 |
(resolve_tac [trans] 1), |
|
266 |
(REPEAT (resolve_tac (prems @ (prems RL [sym])) 1)) ]); |
|
267 |
||
268 |
(** Congruence rules for predicate letters **) |
|
269 |
||
779 | 270 |
qed_goal "pred1_cong" IFOL.thy |
0 | 271 |
"a=a' ==> P(a) <-> P(a')" |
272 |
(fn prems => |
|
273 |
[ (cut_facts_tac prems 1), |
|
274 |
(rtac iffI 1), |
|
275 |
(DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); |
|
276 |
||
779 | 277 |
qed_goal "pred2_cong" IFOL.thy |
0 | 278 |
"[| a=a'; b=b' |] ==> P(a,b) <-> P(a',b')" |
279 |
(fn prems => |
|
280 |
[ (cut_facts_tac prems 1), |
|
281 |
(rtac iffI 1), |
|
282 |
(DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); |
|
283 |
||
779 | 284 |
qed_goal "pred3_cong" IFOL.thy |
0 | 285 |
"[| a=a'; b=b'; c=c' |] ==> P(a,b,c) <-> P(a',b',c')" |
286 |
(fn prems => |
|
287 |
[ (cut_facts_tac prems 1), |
|
288 |
(rtac iffI 1), |
|
289 |
(DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); |
|
290 |
||
291 |
(*special cases for free variables P, Q, R, S -- up to 3 arguments*) |
|
292 |
||
293 |
val pred_congs = |
|
294 |
flat (map (fn c => |
|
295 |
map (fn th => read_instantiate [("P",c)] th) |
|
296 |
[pred1_cong,pred2_cong,pred3_cong]) |
|
297 |
(explode"PQRS")); |
|
298 |
||
299 |
(*special case for the equality predicate!*) |
|
300 |
val eq_cong = read_instantiate [("P","op =")] pred2_cong; |
|
301 |
||
302 |
||
303 |
(*** Simplifications of assumed implications. |
|
304 |
Roy Dyckhoff has proved that conj_impE, disj_impE, and imp_impE |
|
305 |
used with mp_tac (restricted to atomic formulae) is COMPLETE for |
|
306 |
intuitionistic propositional logic. See |
|
307 |
R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic |
|
308 |
(preprint, University of St Andrews, 1991) ***) |
|
309 |
||
779 | 310 |
qed_goal "conj_impE" IFOL.thy |
0 | 311 |
"[| (P&Q)-->S; P-->(Q-->S) ==> R |] ==> R" |
312 |
(fn major::prems=> |
|
313 |
[ (REPEAT (ares_tac ([conjI, impI, major RS mp]@prems) 1)) ]); |
|
314 |
||
779 | 315 |
qed_goal "disj_impE" IFOL.thy |
0 | 316 |
"[| (P|Q)-->S; [| P-->S; Q-->S |] ==> R |] ==> R" |
317 |
(fn major::prems=> |
|
318 |
[ (DEPTH_SOLVE (ares_tac ([disjI1, disjI2, impI, major RS mp]@prems) 1)) ]); |
|
319 |
||
320 |
(*Simplifies the implication. Classical version is stronger. |
|
321 |
Still UNSAFE since Q must be provable -- backtracking needed. *) |
|
779 | 322 |
qed_goal "imp_impE" IFOL.thy |
0 | 323 |
"[| (P-->Q)-->S; [| P; Q-->S |] ==> Q; S ==> R |] ==> R" |
324 |
(fn major::prems=> |
|
325 |
[ (REPEAT (ares_tac ([impI, major RS mp]@prems) 1)) ]); |
|
326 |
||
327 |
(*Simplifies the implication. Classical version is stronger. |
|
328 |
Still UNSAFE since ~P must be provable -- backtracking needed. *) |
|
779 | 329 |
qed_goal "not_impE" IFOL.thy |
0 | 330 |
"[| ~P --> S; P ==> False; S ==> R |] ==> R" |
331 |
(fn major::prems=> |
|
332 |
[ (REPEAT (ares_tac ([notI, impI, major RS mp]@prems) 1)) ]); |
|
333 |
||
334 |
(*Simplifies the implication. UNSAFE. *) |
|
779 | 335 |
qed_goal "iff_impE" IFOL.thy |
0 | 336 |
"[| (P<->Q)-->S; [| P; Q-->S |] ==> Q; [| Q; P-->S |] ==> P; \ |
337 |
\ S ==> R |] ==> R" |
|
338 |
(fn major::prems=> |
|
339 |
[ (REPEAT (ares_tac ([iffI, impI, major RS mp]@prems) 1)) ]); |
|
340 |
||
341 |
(*What if (ALL x.~~P(x)) --> ~~(ALL x.P(x)) is an assumption? UNSAFE*) |
|
779 | 342 |
qed_goal "all_impE" IFOL.thy |
0 | 343 |
"[| (ALL x.P(x))-->S; !!x.P(x); S ==> R |] ==> R" |
344 |
(fn major::prems=> |
|
345 |
[ (REPEAT (ares_tac ([allI, impI, major RS mp]@prems) 1)) ]); |
|
346 |
||
347 |
(*Unsafe: (EX x.P(x))-->S is equivalent to ALL x.P(x)-->S. *) |
|
779 | 348 |
qed_goal "ex_impE" IFOL.thy |
0 | 349 |
"[| (EX x.P(x))-->S; P(x)-->S ==> R |] ==> R" |
350 |
(fn major::prems=> |
|
351 |
[ (REPEAT (ares_tac ([exI, impI, major RS mp]@prems) 1)) ]); |