author | wenzelm |
Tue, 01 Dec 1998 14:48:24 +0100 | |
changeset 6004 | 47705248cf80 |
parent 5888 | d8e51792ca85 |
child 6090 | 78c068b838ff |
permissions | -rw-r--r-- |
1459 | 1 |
(* Title: FOL/IFOL.ML |
0 | 2 |
ID: $Id$ |
1459 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
1280 | 6 |
Tactics and lemmas for IFOL.thy (intuitionistic first-order logic) |
0 | 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 |
3835 | 29 |
"[| ALL x. P(x); P(x) ==> R |] ==> R" |
0 | 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 |
3835 | 34 |
"[| ALL x. P(x); [| P(x); ALL x. P(x) |] ==> R \ |
0 | 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 |
||
1891 | 49 |
qed_goal "rev_notE" IFOL.thy "!!P R. [| P; ~P |] ==> R" |
50 |
(fn _ => [REPEAT (ares_tac [notE] 1)]); |
|
51 |
||
0 | 52 |
(*This is useful with the special implication rules for each kind of P. *) |
779 | 53 |
qed_goal "not_to_imp" IFOL.thy |
0 | 54 |
"[| ~P; (P-->False) ==> Q |] ==> Q" |
55 |
(fn prems=> [ (REPEAT (ares_tac (prems@[impI,notE]) 1)) ]); |
|
56 |
||
1002 | 57 |
(* For substitution into an assumption P, reduce Q to P-->Q, substitute into |
0 | 58 |
this implication, then apply impI to move P back into the assumptions. |
59 |
To specify P use something like |
|
60 |
eres_inst_tac [ ("P","ALL y. ?S(x,y)") ] rev_mp 1 *) |
|
779 | 61 |
qed_goal "rev_mp" IFOL.thy "[| P; P --> Q |] ==> Q" |
0 | 62 |
(fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); |
63 |
||
64 |
(*Contrapositive of an inference rule*) |
|
779 | 65 |
qed_goal "contrapos" IFOL.thy "[| ~Q; P==>Q |] ==> ~P" |
0 | 66 |
(fn [major,minor]=> |
67 |
[ (rtac (major RS notE RS notI) 1), |
|
68 |
(etac minor 1) ]); |
|
69 |
||
70 |
||
71 |
(*** Modus Ponens Tactics ***) |
|
72 |
||
73 |
(*Finds P-->Q and P in the assumptions, replaces implication by Q *) |
|
74 |
fun mp_tac i = eresolve_tac [notE,impE] i THEN assume_tac i; |
|
75 |
||
76 |
(*Like mp_tac but instantiates no variables*) |
|
77 |
fun eq_mp_tac i = eresolve_tac [notE,impE] i THEN eq_assume_tac i; |
|
78 |
||
79 |
||
80 |
(*** If-and-only-if ***) |
|
81 |
||
779 | 82 |
qed_goalw "iffI" IFOL.thy [iff_def] |
0 | 83 |
"[| P ==> Q; Q ==> P |] ==> P<->Q" |
84 |
(fn prems=> [ (REPEAT (ares_tac (prems@[conjI, impI]) 1)) ]); |
|
85 |
||
86 |
||
87 |
(*Observe use of rewrite_rule to unfold "<->" in meta-assumptions (prems) *) |
|
779 | 88 |
qed_goalw "iffE" IFOL.thy [iff_def] |
0 | 89 |
"[| P <-> Q; [| P-->Q; Q-->P |] ==> R |] ==> R" |
1459 | 90 |
(fn prems => [ (rtac conjE 1), (REPEAT (ares_tac prems 1)) ]); |
0 | 91 |
|
92 |
(* Destruct rules for <-> similar to Modus Ponens *) |
|
93 |
||
779 | 94 |
qed_goalw "iffD1" IFOL.thy [iff_def] "[| P <-> Q; P |] ==> Q" |
0 | 95 |
(fn prems => [ (rtac (conjunct1 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); |
96 |
||
779 | 97 |
qed_goalw "iffD2" IFOL.thy [iff_def] "[| P <-> Q; Q |] ==> P" |
0 | 98 |
(fn prems => [ (rtac (conjunct2 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); |
99 |
||
4462 | 100 |
qed_goal "rev_iffD1" IFOL.thy "!!P. [| P; P <-> Q |] ==> Q" |
101 |
(fn _ => [etac iffD1 1, assume_tac 1]); |
|
102 |
||
103 |
qed_goal "rev_iffD2" IFOL.thy "!!P. [| Q; P <-> Q |] ==> P" |
|
104 |
(fn _ => [etac iffD2 1, assume_tac 1]); |
|
105 |
||
779 | 106 |
qed_goal "iff_refl" IFOL.thy "P <-> P" |
0 | 107 |
(fn _ => [ (REPEAT (ares_tac [iffI] 1)) ]); |
108 |
||
779 | 109 |
qed_goal "iff_sym" IFOL.thy "Q <-> P ==> P <-> Q" |
0 | 110 |
(fn [major] => |
111 |
[ (rtac (major RS iffE) 1), |
|
112 |
(rtac iffI 1), |
|
113 |
(REPEAT (eresolve_tac [asm_rl,mp] 1)) ]); |
|
114 |
||
779 | 115 |
qed_goal "iff_trans" IFOL.thy |
0 | 116 |
"!!P Q R. [| P <-> Q; Q<-> R |] ==> P <-> R" |
117 |
(fn _ => |
|
118 |
[ (rtac iffI 1), |
|
119 |
(REPEAT (eresolve_tac [asm_rl,iffE] 1 ORELSE mp_tac 1)) ]); |
|
120 |
||
121 |
||
122 |
(*** Unique existence. NOTE THAT the following 2 quantifications |
|
123 |
EX!x such that [EX!y such that P(x,y)] (sequential) |
|
124 |
EX!x,y such that P(x,y) (simultaneous) |
|
125 |
do NOT mean the same thing. The parser treats EX!x y.P(x,y) as sequential. |
|
126 |
***) |
|
127 |
||
779 | 128 |
qed_goalw "ex1I" IFOL.thy [ex1_def] |
0 | 129 |
"[| P(a); !!x. P(x) ==> x=a |] ==> EX! x. P(x)" |
130 |
(fn prems => [ (REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)) ]); |
|
131 |
||
2843 | 132 |
(*Sometimes easier to use: the premises have no shared variables. Safe!*) |
779 | 133 |
qed_goal "ex_ex1I" IFOL.thy |
3835 | 134 |
"[| EX x. P(x); !!x y. [| P(x); P(y) |] ==> x=y |] ==> EX! x. P(x)" |
12
f17d542276b6
Added ex_ex1I: new introduction rule for the EX! quantifier.
lcp
parents:
0
diff
changeset
|
135 |
(fn [ex,eq] => [ (rtac (ex RS exE) 1), |
1459 | 136 |
(REPEAT (ares_tac [ex1I,eq] 1)) ]); |
12
f17d542276b6
Added ex_ex1I: new introduction rule for the EX! quantifier.
lcp
parents:
0
diff
changeset
|
137 |
|
779 | 138 |
qed_goalw "ex1E" IFOL.thy [ex1_def] |
3835 | 139 |
"[| EX! x. P(x); !!x. [| P(x); ALL y. P(y) --> y=x |] ==> R |] ==> R" |
0 | 140 |
(fn prems => |
141 |
[ (cut_facts_tac prems 1), |
|
142 |
(REPEAT (eresolve_tac [exE,conjE] 1 ORELSE ares_tac prems 1)) ]); |
|
143 |
||
144 |
||
145 |
(*** <-> congruence rules for simplification ***) |
|
146 |
||
147 |
(*Use iffE on a premise. For conj_cong, imp_cong, all_cong, ex_cong*) |
|
148 |
fun iff_tac prems i = |
|
149 |
resolve_tac (prems RL [iffE]) i THEN |
|
150 |
REPEAT1 (eresolve_tac [asm_rl,mp] i); |
|
151 |
||
779 | 152 |
qed_goal "conj_cong" IFOL.thy |
0 | 153 |
"[| P <-> P'; P' ==> Q <-> Q' |] ==> (P&Q) <-> (P'&Q')" |
154 |
(fn prems => |
|
155 |
[ (cut_facts_tac prems 1), |
|
156 |
(REPEAT (ares_tac [iffI,conjI] 1 |
|
157 |
ORELSE eresolve_tac [iffE,conjE,mp] 1 |
|
158 |
ORELSE iff_tac prems 1)) ]); |
|
159 |
||
793 | 160 |
(*Reversed congruence rule! Used in ZF/Order*) |
161 |
qed_goal "conj_cong2" IFOL.thy |
|
162 |
"[| P <-> P'; P' ==> Q <-> Q' |] ==> (Q&P) <-> (Q'&P')" |
|
163 |
(fn prems => |
|
164 |
[ (cut_facts_tac prems 1), |
|
165 |
(REPEAT (ares_tac [iffI,conjI] 1 |
|
166 |
ORELSE eresolve_tac [iffE,conjE,mp] 1 |
|
167 |
ORELSE iff_tac prems 1)) ]); |
|
168 |
||
779 | 169 |
qed_goal "disj_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,disjE,disjI1,disjI2] 1 |
|
174 |
ORELSE ares_tac [iffI] 1 |
|
175 |
ORELSE mp_tac 1)) ]); |
|
176 |
||
779 | 177 |
qed_goal "imp_cong" IFOL.thy |
0 | 178 |
"[| P <-> P'; P' ==> Q <-> Q' |] ==> (P-->Q) <-> (P'-->Q')" |
179 |
(fn prems => |
|
180 |
[ (cut_facts_tac prems 1), |
|
181 |
(REPEAT (ares_tac [iffI,impI] 1 |
|
1459 | 182 |
ORELSE etac iffE 1 |
0 | 183 |
ORELSE mp_tac 1 ORELSE iff_tac prems 1)) ]); |
184 |
||
779 | 185 |
qed_goal "iff_cong" IFOL.thy |
0 | 186 |
"[| P <-> P'; Q <-> Q' |] ==> (P<->Q) <-> (P'<->Q')" |
187 |
(fn prems => |
|
188 |
[ (cut_facts_tac prems 1), |
|
1459 | 189 |
(REPEAT (etac iffE 1 |
0 | 190 |
ORELSE ares_tac [iffI] 1 |
191 |
ORELSE mp_tac 1)) ]); |
|
192 |
||
779 | 193 |
qed_goal "not_cong" IFOL.thy |
0 | 194 |
"P <-> P' ==> ~P <-> ~P'" |
195 |
(fn prems => |
|
196 |
[ (cut_facts_tac prems 1), |
|
197 |
(REPEAT (ares_tac [iffI,notI] 1 |
|
198 |
ORELSE mp_tac 1 |
|
199 |
ORELSE eresolve_tac [iffE,notE] 1)) ]); |
|
200 |
||
779 | 201 |
qed_goal "all_cong" IFOL.thy |
3835 | 202 |
"(!!x. P(x) <-> Q(x)) ==> (ALL x. P(x)) <-> (ALL x. Q(x))" |
0 | 203 |
(fn prems => |
204 |
[ (REPEAT (ares_tac [iffI,allI] 1 |
|
205 |
ORELSE mp_tac 1 |
|
1459 | 206 |
ORELSE etac allE 1 ORELSE iff_tac prems 1)) ]); |
0 | 207 |
|
779 | 208 |
qed_goal "ex_cong" IFOL.thy |
3835 | 209 |
"(!!x. P(x) <-> Q(x)) ==> (EX x. P(x)) <-> (EX x. Q(x))" |
0 | 210 |
(fn prems => |
1459 | 211 |
[ (REPEAT (etac exE 1 ORELSE ares_tac [iffI,exI] 1 |
0 | 212 |
ORELSE mp_tac 1 |
213 |
ORELSE iff_tac prems 1)) ]); |
|
214 |
||
779 | 215 |
qed_goal "ex1_cong" IFOL.thy |
3835 | 216 |
"(!!x. P(x) <-> Q(x)) ==> (EX! x. P(x)) <-> (EX! x. Q(x))" |
0 | 217 |
(fn prems => |
218 |
[ (REPEAT (eresolve_tac [ex1E, spec RS mp] 1 ORELSE ares_tac [iffI,ex1I] 1 |
|
219 |
ORELSE mp_tac 1 |
|
220 |
ORELSE iff_tac prems 1)) ]); |
|
221 |
||
222 |
(*** Equality rules ***) |
|
223 |
||
779 | 224 |
qed_goal "sym" IFOL.thy "a=b ==> b=a" |
0 | 225 |
(fn [major] => [ (rtac (major RS subst) 1), (rtac refl 1) ]); |
226 |
||
779 | 227 |
qed_goal "trans" IFOL.thy "[| a=b; b=c |] ==> a=c" |
0 | 228 |
(fn [prem1,prem2] => [ (rtac (prem2 RS subst) 1), (rtac prem1 1) ]); |
229 |
||
230 |
(** ~ b=a ==> ~ a=b **) |
|
231 |
val [not_sym] = compose(sym,2,contrapos); |
|
232 |
||
5309
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
233 |
|
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
234 |
(* Two theorms for rewriting only one instance of a definition: |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
235 |
the first for definitions of formulae and the second for terms *) |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
236 |
|
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
237 |
val prems = goal IFOL.thy "(A == B) ==> A <-> B"; |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
238 |
by (rewrite_goals_tac prems); |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
239 |
by (rtac iff_refl 1); |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
240 |
qed "def_imp_iff"; |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
241 |
|
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
242 |
val prems = goal IFOL.thy "(A == B) ==> A = B"; |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
243 |
by (rewrite_goals_tac prems); |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
244 |
by (rtac refl 1); |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
245 |
qed "def_imp_eq"; |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
246 |
|
2037 | 247 |
(*Substitution: rule and tactic*) |
248 |
bind_thm ("ssubst", sym RS subst); |
|
5139
013ea0f023e3
stac now uses CHANGED_GOAL and correctly fails when it has no useful effect,
paulson
parents:
4462
diff
changeset
|
249 |
|
5309
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
250 |
(*Apply an equality or definition ONCE. |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
251 |
Fails unless the substitution has an effect*) |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
252 |
fun stac th = |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
253 |
let val th' = th RS def_imp_eq handle _ => th |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
254 |
in CHANGED_GOAL (rtac (th' RS ssubst)) |
01c2b317da88
stac now handles definitions as well as equalities
paulson
parents:
5139
diff
changeset
|
255 |
end; |
0 | 256 |
|
257 |
(*A special case of ex1E that would otherwise need quantifier expansion*) |
|
779 | 258 |
qed_goal "ex1_equalsE" IFOL.thy |
3835 | 259 |
"[| EX! x. P(x); P(a); P(b) |] ==> a=b" |
0 | 260 |
(fn prems => |
261 |
[ (cut_facts_tac prems 1), |
|
262 |
(etac ex1E 1), |
|
263 |
(rtac trans 1), |
|
264 |
(rtac sym 2), |
|
265 |
(REPEAT (eresolve_tac [asm_rl, spec RS mp] 1)) ]); |
|
266 |
||
267 |
(** Polymorphic congruence rules **) |
|
268 |
||
779 | 269 |
qed_goal "subst_context" IFOL.thy |
0 | 270 |
"[| a=b |] ==> t(a)=t(b)" |
271 |
(fn prems=> |
|
272 |
[ (resolve_tac (prems RL [ssubst]) 1), |
|
1459 | 273 |
(rtac refl 1) ]); |
0 | 274 |
|
779 | 275 |
qed_goal "subst_context2" IFOL.thy |
0 | 276 |
"[| a=b; c=d |] ==> t(a,c)=t(b,d)" |
277 |
(fn prems=> |
|
278 |
[ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); |
|
279 |
||
779 | 280 |
qed_goal "subst_context3" IFOL.thy |
0 | 281 |
"[| a=b; c=d; e=f |] ==> t(a,c,e)=t(b,d,f)" |
282 |
(fn prems=> |
|
283 |
[ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); |
|
284 |
||
285 |
(*Useful with eresolve_tac for proving equalties from known equalities. |
|
1459 | 286 |
a = b |
287 |
| | |
|
288 |
c = d *) |
|
779 | 289 |
qed_goal "box_equals" IFOL.thy |
0 | 290 |
"[| a=b; a=c; b=d |] ==> c=d" |
291 |
(fn prems=> |
|
1459 | 292 |
[ (rtac trans 1), |
293 |
(rtac trans 1), |
|
294 |
(rtac sym 1), |
|
0 | 295 |
(REPEAT (resolve_tac prems 1)) ]); |
296 |
||
297 |
(*Dual of box_equals: for proving equalities backwards*) |
|
779 | 298 |
qed_goal "simp_equals" IFOL.thy |
0 | 299 |
"[| a=c; b=d; c=d |] ==> a=b" |
300 |
(fn prems=> |
|
1459 | 301 |
[ (rtac trans 1), |
302 |
(rtac trans 1), |
|
0 | 303 |
(REPEAT (resolve_tac (prems @ (prems RL [sym])) 1)) ]); |
304 |
||
305 |
(** Congruence rules for predicate letters **) |
|
306 |
||
779 | 307 |
qed_goal "pred1_cong" IFOL.thy |
0 | 308 |
"a=a' ==> P(a) <-> P(a')" |
309 |
(fn prems => |
|
310 |
[ (cut_facts_tac prems 1), |
|
311 |
(rtac iffI 1), |
|
312 |
(DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); |
|
313 |
||
779 | 314 |
qed_goal "pred2_cong" IFOL.thy |
0 | 315 |
"[| a=a'; b=b' |] ==> P(a,b) <-> P(a',b')" |
316 |
(fn prems => |
|
317 |
[ (cut_facts_tac prems 1), |
|
318 |
(rtac iffI 1), |
|
319 |
(DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); |
|
320 |
||
779 | 321 |
qed_goal "pred3_cong" IFOL.thy |
0 | 322 |
"[| a=a'; b=b'; c=c' |] ==> P(a,b,c) <-> P(a',b',c')" |
323 |
(fn prems => |
|
324 |
[ (cut_facts_tac prems 1), |
|
325 |
(rtac iffI 1), |
|
326 |
(DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); |
|
327 |
||
328 |
(*special cases for free variables P, Q, R, S -- up to 3 arguments*) |
|
329 |
||
330 |
val pred_congs = |
|
331 |
flat (map (fn c => |
|
1459 | 332 |
map (fn th => read_instantiate [("P",c)] th) |
333 |
[pred1_cong,pred2_cong,pred3_cong]) |
|
334 |
(explode"PQRS")); |
|
0 | 335 |
|
336 |
(*special case for the equality predicate!*) |
|
337 |
val eq_cong = read_instantiate [("P","op =")] pred2_cong; |
|
338 |
||
339 |
||
340 |
(*** Simplifications of assumed implications. |
|
341 |
Roy Dyckhoff has proved that conj_impE, disj_impE, and imp_impE |
|
342 |
used with mp_tac (restricted to atomic formulae) is COMPLETE for |
|
343 |
intuitionistic propositional logic. See |
|
344 |
R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic |
|
345 |
(preprint, University of St Andrews, 1991) ***) |
|
346 |
||
779 | 347 |
qed_goal "conj_impE" IFOL.thy |
0 | 348 |
"[| (P&Q)-->S; P-->(Q-->S) ==> R |] ==> R" |
349 |
(fn major::prems=> |
|
350 |
[ (REPEAT (ares_tac ([conjI, impI, major RS mp]@prems) 1)) ]); |
|
351 |
||
779 | 352 |
qed_goal "disj_impE" IFOL.thy |
0 | 353 |
"[| (P|Q)-->S; [| P-->S; Q-->S |] ==> R |] ==> R" |
354 |
(fn major::prems=> |
|
355 |
[ (DEPTH_SOLVE (ares_tac ([disjI1, disjI2, impI, major RS mp]@prems) 1)) ]); |
|
356 |
||
357 |
(*Simplifies the implication. Classical version is stronger. |
|
358 |
Still UNSAFE since Q must be provable -- backtracking needed. *) |
|
779 | 359 |
qed_goal "imp_impE" IFOL.thy |
0 | 360 |
"[| (P-->Q)-->S; [| P; Q-->S |] ==> Q; S ==> R |] ==> R" |
361 |
(fn major::prems=> |
|
362 |
[ (REPEAT (ares_tac ([impI, major RS mp]@prems) 1)) ]); |
|
363 |
||
364 |
(*Simplifies the implication. Classical version is stronger. |
|
365 |
Still UNSAFE since ~P must be provable -- backtracking needed. *) |
|
779 | 366 |
qed_goal "not_impE" IFOL.thy |
0 | 367 |
"[| ~P --> S; P ==> False; S ==> R |] ==> R" |
368 |
(fn major::prems=> |
|
369 |
[ (REPEAT (ares_tac ([notI, impI, major RS mp]@prems) 1)) ]); |
|
370 |
||
371 |
(*Simplifies the implication. UNSAFE. *) |
|
779 | 372 |
qed_goal "iff_impE" IFOL.thy |
0 | 373 |
"[| (P<->Q)-->S; [| P; Q-->S |] ==> Q; [| Q; P-->S |] ==> P; \ |
374 |
\ S ==> R |] ==> R" |
|
375 |
(fn major::prems=> |
|
376 |
[ (REPEAT (ares_tac ([iffI, impI, major RS mp]@prems) 1)) ]); |
|
377 |
||
378 |
(*What if (ALL x.~~P(x)) --> ~~(ALL x.P(x)) is an assumption? UNSAFE*) |
|
779 | 379 |
qed_goal "all_impE" IFOL.thy |
3835 | 380 |
"[| (ALL x. P(x))-->S; !!x. P(x); S ==> R |] ==> R" |
0 | 381 |
(fn major::prems=> |
382 |
[ (REPEAT (ares_tac ([allI, impI, major RS mp]@prems) 1)) ]); |
|
383 |
||
384 |
(*Unsafe: (EX x.P(x))-->S is equivalent to ALL x.P(x)-->S. *) |
|
779 | 385 |
qed_goal "ex_impE" IFOL.thy |
3835 | 386 |
"[| (EX x. P(x))-->S; P(x)-->S ==> R |] ==> R" |
0 | 387 |
(fn major::prems=> |
388 |
[ (REPEAT (ares_tac ([exI, impI, major RS mp]@prems) 1)) ]); |
|
821 | 389 |
|
1608 | 390 |
(*** Courtesy of Krzysztof Grabczewski ***) |
391 |
||
821 | 392 |
val major::prems = goal IFOL.thy "[| P|Q; P==>R; Q==>S |] ==> R|S"; |
1459 | 393 |
by (rtac (major RS disjE) 1); |
821 | 394 |
by (REPEAT (eresolve_tac (prems RL [disjI1, disjI2]) 1)); |
395 |
qed "disj_imp_disj"; |
|
1608 | 396 |
|
3722
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
397 |
|
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
398 |
(** strip ALL and --> from proved goal while preserving ALL-bound var names **) |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
399 |
|
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
400 |
local |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
401 |
|
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
402 |
(* Use XXX to avoid forall_intr failing because of duplicate variable name *) |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
403 |
val myspec = read_instantiate [("P","?XXX")] spec; |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
404 |
val _ $ (_ $ (vx as Var(_,vxT))) = concl_of myspec; |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
405 |
val cvx = cterm_of (#sign(rep_thm myspec)) vx; |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
406 |
val aspec = forall_intr cvx myspec; |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
407 |
|
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
408 |
in |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
409 |
|
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
410 |
fun RSspec th = |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
411 |
(case concl_of th of |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
412 |
_ $ (Const("All",_) $ Abs(a,_,_)) => |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
413 |
let val ca = cterm_of (#sign(rep_thm th)) (Var((a,0),vxT)) |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
414 |
in th RS forall_elim ca aspec end |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
415 |
| _ => raise THM("RSspec",0,[th])); |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
416 |
|
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
417 |
fun RSmp th = |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
418 |
(case concl_of th of |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
419 |
_ $ (Const("op -->",_)$_$_) => th RS mp |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
420 |
| _ => raise THM("RSmp",0,[th])); |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
421 |
|
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
422 |
fun normalize_thm funs = |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
423 |
let fun trans [] th = th |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
424 |
| trans (f::fs) th = (trans funs (f th)) handle THM _ => trans fs th |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
425 |
in trans funs end; |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
426 |
|
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
427 |
fun qed_spec_mp name = |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
428 |
let val thm = normalize_thm [RSspec,RSmp] (result()) |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
429 |
in bind_thm(name, thm) end; |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
430 |
|
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
431 |
fun qed_goal_spec_mp name thy s p = |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
432 |
bind_thm (name, normalize_thm [RSspec,RSmp] (prove_goal thy s p)); |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
433 |
|
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
434 |
fun qed_goalw_spec_mp name thy defs s p = |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
435 |
bind_thm (name, normalize_thm [RSspec,RSmp] (prove_goalw thy defs s p)); |
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
436 |
|
24af9e73451e
qed_spec_mp, normalize_thm, etc. copied in from HOL.ML
paulson
parents:
2843
diff
changeset
|
437 |
end; |
5888 | 438 |
|
439 |
||
440 |
(* attributes *) |
|
441 |
||
442 |
local |
|
443 |
||
444 |
fun gen_rulify x = Attrib.no_args (Attribute.rule (K (normalize_thm [RSspec, RSmp]))) x; |
|
445 |
||
446 |
in |
|
447 |
||
448 |
val attrib_setup = |
|
449 |
[Attrib.add_attributes |
|
450 |
[("rulify", (gen_rulify, gen_rulify), "put theorem into standard rule form")]]; |
|
451 |
||
452 |
end; |