| author | Christian Urban <urbanc@in.tum.de> |
| Thu, 05 Feb 2009 22:44:11 +0000 | |
| changeset 100 | dd8eebae11ec |
| parent 99 | de625e5f6a36 |
| child 102 | 5e309df58557 |
| permissions | -rw-r--r-- |
| 93 | 1 |
theory Tactical |
| 99 | 2 |
imports Base FirstSteps |
| 93 | 3 |
begin |
4 |
||
5 |
chapter {* Tactical Reasoning\label{chp:tactical} *}
|
|
6 |
||
7 |
text {*
|
|
8 |
||
9 |
The main reason for descending to the ML-level of Isabelle is to be able to |
|
10 |
implement automatic proof procedures. Such proof procedures usually lessen |
|
11 |
considerably the burden of manual reasoning, for example, when introducing |
|
12 |
new definitions. These proof procedures are centred around refining a goal |
|
13 |
state using tactics. This is similar to the @{text apply}-style reasoning at
|
|
14 |
the user level, where goals are modified in a sequence of proof steps until |
|
| 99 | 15 |
all of them are solved. However, there are also more structured operations |
16 |
that help with handling of variables and assumptions. |
|
| 93 | 17 |
*} |
18 |
||
19 |
section {* Tactical Reasoning *}
|
|
20 |
||
21 |
text {*
|
|
22 |
To see how tactics work, let us first transcribe a simple @{text apply}-style proof
|
|
23 |
into ML. Consider the following proof. |
|
24 |
*} |
|
25 |
||
26 |
lemma disj_swap: "P \<or> Q \<Longrightarrow> Q \<or> P" |
|
27 |
apply(erule disjE) |
|
28 |
apply(rule disjI2) |
|
29 |
apply(assumption) |
|
30 |
apply(rule disjI1) |
|
31 |
apply(assumption) |
|
32 |
done |
|
33 |
||
34 |
text {*
|
|
35 |
This proof translates to the following ML-code. |
|
36 |
||
37 |
@{ML_response_fake [display,gray]
|
|
38 |
"let |
|
39 |
val ctxt = @{context}
|
|
40 |
val goal = @{prop \"P \<or> Q \<Longrightarrow> Q \<or> P\"}
|
|
41 |
in |
|
42 |
Goal.prove ctxt [\"P\", \"Q\"] [] goal |
|
43 |
(fn _ => |
|
44 |
etac @{thm disjE} 1
|
|
45 |
THEN rtac @{thm disjI2} 1
|
|
46 |
THEN atac 1 |
|
47 |
THEN rtac @{thm disjI1} 1
|
|
48 |
THEN atac 1) |
|
49 |
end" "?P \<or> ?Q \<Longrightarrow> ?Q \<or> ?P"} |
|
50 |
||
51 |
To start the proof, the function @{ML "Goal.prove"}~@{text "ctxt xs As C
|
|
| 99 | 52 |
tac"} sets up a goal state for proving the goal @{text C}
|
53 |
(that is @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"} in the proof at hand) under the
|
|
54 |
assumptions @{text As} (happens to be empty) with the variables
|
|
| 93 | 55 |
@{text xs} that will be generalised once the goal is proved (in our case
|
56 |
@{text P} and @{text Q}). The @{text "tac"} is the tactic that proves the goal;
|
|
57 |
it can make use of the local assumptions (there are none in this example). |
|
58 |
The functions @{ML etac}, @{ML rtac} and @{ML atac} correspond to
|
|
59 |
@{text erule}, @{text rule} and @{text assumption}, respectively.
|
|
| 99 | 60 |
The operator @{ML THEN} strings the tactics together.
|
| 93 | 61 |
|
62 |
\begin{readmore}
|
|
| 99 | 63 |
To learn more about the function @{ML Goal.prove} see \isccite{sec:results}
|
64 |
and the file @{ML_file "Pure/goal.ML"}. For more information about the
|
|
65 |
internals of goals see \isccite{sec:tactical-goals}. See @{ML_file
|
|
66 |
"Pure/tactic.ML"} and @{ML_file "Pure/tctical.ML"} for the code of basic
|
|
67 |
tactics and tactic combinators; see also Chapters 3 and 4 in the old |
|
68 |
Isabelle Reference Manual. |
|
| 93 | 69 |
\end{readmore}
|
70 |
||
71 |
Note that we used antiquotations for referencing the theorems. We could also |
|
72 |
just have written @{ML "etac disjE 1"} and so on, but this is considered bad
|
|
73 |
style. The reason is that the binding for @{ML disjE} can be re-assigned by
|
|
74 |
the user and thus one does not have complete control over which theorem is |
|
75 |
actually applied. This problem is nicely prevented by using antiquotations, |
|
76 |
because then the theorems are fixed statically at compile-time. |
|
77 |
||
78 |
During the development of automatic proof procedures, it will often be necessary |
|
79 |
to test a tactic on examples. This can be conveniently |
|
80 |
done with the command \isacommand{apply}@{text "(tactic \<verbopen> \<dots> \<verbclose>)"}.
|
|
81 |
Consider the following sequence of tactics |
|
82 |
*} |
|
83 |
||
84 |
ML{*val foo_tac =
|
|
85 |
(etac @{thm disjE} 1
|
|
86 |
THEN rtac @{thm disjI2} 1
|
|
87 |
THEN atac 1 |
|
88 |
THEN rtac @{thm disjI1} 1
|
|
89 |
THEN atac 1)*} |
|
90 |
||
91 |
text {* and the Isabelle proof: *}
|
|
92 |
||
93 |
lemma "P \<or> Q \<Longrightarrow> Q \<or> P" |
|
94 |
apply(tactic {* foo_tac *})
|
|
95 |
done |
|
96 |
||
97 |
text {*
|
|
| 99 | 98 |
By using @{text "tactic \<verbopen> \<dots> \<verbclose>"} in the apply-step,
|
99 |
you can call @{ML foo_tac} or any function that returns a tactic from the
|
|
100 |
user level of Isabelle. |
|
| 93 | 101 |
|
102 |
As can be seen, each tactic in @{ML foo_tac} has a hard-coded number that
|
|
103 |
stands for the subgoal analysed by the tactic. In our case, we only focus on the first |
|
104 |
subgoal. This is sometimes wanted, but usually not. To avoid the explicit numbering in the |
|
105 |
tactic, you can write |
|
106 |
*} |
|
107 |
||
108 |
ML{*val foo_tac' =
|
|
109 |
(etac @{thm disjE}
|
|
110 |
THEN' rtac @{thm disjI2}
|
|
111 |
THEN' atac |
|
112 |
THEN' rtac @{thm disjI1}
|
|
113 |
THEN' atac)*} |
|
114 |
||
115 |
text {*
|
|
116 |
and then give the number for the subgoal explicitly when the tactic is |
|
| 99 | 117 |
called. For every operator that combines tactics, such a primed version |
118 |
exists. So in the next proof we can first discharge the second subgoal, |
|
| 93 | 119 |
and after that the first. |
120 |
*} |
|
121 |
||
122 |
lemma "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1" |
|
123 |
and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2" |
|
124 |
apply(tactic {* foo_tac' 2 *})
|
|
125 |
apply(tactic {* foo_tac' 1 *})
|
|
126 |
done |
|
127 |
||
128 |
text {*
|
|
| 99 | 129 |
(FIXME: maybe add something about how each goal state is interpreted |
130 |
as a theorem) |
|
131 |
||
132 |
The tactic @{ML foo_tac} (and @{ML foo_tac'}) are very specific for
|
|
133 |
analysing goals only of the form @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. If the goal is not
|
|
134 |
of this form, then @{ML foo_tac} throws the error message:
|
|
135 |
||
136 |
\begin{isabelle}
|
|
137 |
@{text "*** empty result sequence -- proof command failed"}\\
|
|
138 |
@{text "*** At command \"apply\"."}
|
|
139 |
\end{isabelle}
|
|
140 |
||
141 |
Meaning the tactic failed. The reason for this error message is that tactics |
|
142 |
are functions that map a goal state to a (lazy) sequence of successor states, |
|
143 |
hence the type of a tactic is |
|
| 93 | 144 |
|
145 |
@{text [display, gray] "type tactic = thm -> thm Seq.seq"}
|
|
146 |
||
| 99 | 147 |
It is custom that if a tactic fails, it should return the empty sequence: |
148 |
tactics should not raise exceptions willy-nilly. |
|
| 93 | 149 |
|
| 99 | 150 |
The lazy list of possible successor states shows through to the user-level |
151 |
of Isabelle when using the command \isacommand{back}. For instance in the
|
|
152 |
following proof, there are two possibilities for how to apply |
|
153 |
@{ML foo_tac'}.
|
|
| 93 | 154 |
*} |
155 |
||
156 |
lemma "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P" |
|
157 |
apply(tactic {* foo_tac' 1 *})
|
|
158 |
back |
|
159 |
done |
|
160 |
||
161 |
text {*
|
|
| 99 | 162 |
By using \isacommand{back}, we construct the proof that uses the
|
163 |
second assumption to complete the proof. In more interesting |
|
164 |
situations, different possibilities can lead to different proofs. |
|
165 |
||
| 93 | 166 |
\begin{readmore}
|
167 |
See @{ML_file "Pure/General/seq.ML"} for the implementation of lazy
|
|
168 |
sequences. However in day-to-day Isabelle programming, one rarely |
|
169 |
constructs sequences explicitly, but uses the predefined functions |
|
| 99 | 170 |
instead. |
| 93 | 171 |
\end{readmore}
|
172 |
||
173 |
*} |
|
174 |
||
175 |
section {* Basic Tactics *}
|
|
176 |
||
| 99 | 177 |
text {*
|
178 |
As seen above, the function @{ML atac} corresponds to the assumption tactic.
|
|
179 |
*} |
|
180 |
||
181 |
lemma shows "P \<Longrightarrow> P" |
|
| 93 | 182 |
apply(tactic {* atac 1 *})
|
183 |
done |
|
184 |
||
| 99 | 185 |
text {*
|
186 |
Similarly, @{ML rtac}, @{ML dtac}, @{ML etac} and @{ML ftac} correspond
|
|
187 |
to @{ML_text rule}, @{ML_text drule}, @{ML_text erule} and @{ML_text frule},
|
|
188 |
respectively. Below are three examples with the resulting goal state. How |
|
189 |
they work should therefore be self-explanatory. |
|
190 |
*} |
|
191 |
||
192 |
lemma shows "P \<and> Q" |
|
| 93 | 193 |
apply(tactic {* rtac @{thm conjI} 1 *})
|
| 99 | 194 |
txt{*@{subgoals [display]}*}
|
| 93 | 195 |
(*<*)oops(*>*) |
196 |
||
| 99 | 197 |
lemma shows "P \<and> Q \<Longrightarrow> False" |
| 93 | 198 |
apply(tactic {* etac @{thm conjE} 1 *})
|
| 99 | 199 |
txt{*@{subgoals [display]}*}
|
| 93 | 200 |
(*<*)oops(*>*) |
201 |
||
202 |
lemma shows "False \<and> True \<Longrightarrow> False" |
|
203 |
apply(tactic {* dtac @{thm conjunct2} 1 *})
|
|
| 99 | 204 |
txt{*@{subgoals [display]}*}
|
| 93 | 205 |
(*<*)oops(*>*) |
206 |
||
207 |
text {*
|
|
| 99 | 208 |
As mentioned above, most basic tactics take a number as argument, which |
209 |
addresses to subgoal they are analysing. |
|
210 |
*} |
|
211 |
||
212 |
lemma shows "Foo" and "P \<and> Q" |
|
213 |
apply(tactic {* rtac @{thm conjI} 2 *})
|
|
214 |
txt {*@{subgoals [display]}*}
|
|
215 |
(*<*)oops(*>*) |
|
216 |
||
217 |
text {*
|
|
218 |
Corresponding to @{ML rtac}, there is also the tactic @{ML resolve_tac}, which
|
|
219 |
however expects a list of theorems as arguments. From this list it will apply with |
|
220 |
the first applicable theorem (later theorems that are also applicable can be |
|
221 |
explored via the lazy sequences mechanism). Given the abbreviation |
|
| 93 | 222 |
*} |
223 |
||
| 99 | 224 |
ML{*val resolve_tac_xmp = resolve_tac [@{thm impI}, @{thm conjI}]*}
|
225 |
||
226 |
text {*
|
|
227 |
an example for @{ML resolve_tac} is the following proof where first an outermost
|
|
228 |
implication is analysed and then an outermost conjunction. |
|
229 |
*} |
|
230 |
||
231 |
lemma shows "C \<longrightarrow> (A \<and> B)" and "(A \<longrightarrow> B) \<and> C" |
|
232 |
apply(tactic {* resolve_tac_xmp 1 *})
|
|
233 |
apply(tactic {* resolve_tac_xmp 2 *})
|
|
234 |
txt{* @{subgoals [display]} *}
|
|
235 |
(*<*)oops(*>*) |
|
236 |
||
237 |
text {*
|
|
238 |
Similarly versions exists for @{ML atac} (@{ML assume_tac}), @{ML etac}
|
|
239 |
(@{ML eresolve_tac}) and so on.
|
|
240 |
||
241 |
The tactic @{ML print_tac} is useful for low-level debugging of tactics: it
|
|
242 |
prints out a message and the current goal state. |
|
243 |
*} |
|
244 |
||
245 |
lemma shows "False \<Longrightarrow> True" |
|
| 93 | 246 |
apply(tactic {* print_tac "foo message" *})
|
247 |
(*<*)oops(*>*) |
|
248 |
||
| 99 | 249 |
text {*
|
250 |
Also for useful for debugging, but not exclusively, are the tactics @{ML all_tac} and
|
|
251 |
@{ML no_tac}. The former always succeeds (but does not change the goal state), and
|
|
252 |
the latter always fails. |
|
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
253 |
|
| 99 | 254 |
(FIXME explain RS MRS) |
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
255 |
|
| 99 | 256 |
Often proofs involve elaborate operations on assumptions and variables |
257 |
@{text "\<And>"}-quantified in the goal state. To do such operations on the ML-level
|
|
258 |
using the basic tactics, is very unwieldy and brittle. Some convenience and |
|
259 |
safety is provided by the tactic @{ML SUBPROOF}. This tactic fixes the parameters
|
|
260 |
and binds the various components of a proof state into a record. |
|
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
261 |
*} |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
262 |
|
| 99 | 263 |
text_raw{*
|
264 |
\begin{figure}
|
|
265 |
\begin{isabelle}
|
|
266 |
*} |
|
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
267 |
ML{*fun sp_tac {prems, params, asms, concl, context, schematics} =
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
268 |
let |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
269 |
val str_of_params = str_of_cterms context params |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
270 |
val str_of_asms = str_of_cterms context asms |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
271 |
val str_of_concl = str_of_cterm context concl |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
272 |
val str_of_prems = str_of_thms context prems |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
273 |
val str_of_schms = str_of_cterms context (snd schematics) |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
274 |
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
275 |
val _ = (warning ("params: " ^ str_of_params);
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
276 |
warning ("schematics: " ^ str_of_schms);
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
277 |
warning ("asms: " ^ str_of_asms);
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
278 |
warning ("concl: " ^ str_of_concl);
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
279 |
warning ("prems: " ^ str_of_prems))
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
280 |
in |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
281 |
no_tac |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
282 |
end*} |
| 99 | 283 |
text_raw{*
|
284 |
\end{isabelle}
|
|
285 |
\caption{A function that prints out the various parameters provided by the tactic
|
|
286 |
@{ML SUBPROOF}. It uses the functions extracting strings from @{ML_type cterm}s
|
|
287 |
and @{ML_type thm}s, which are defined in Section~\ref{sec:printing}.\label{fig:sptac}}
|
|
288 |
\end{figure}
|
|
289 |
*} |
|
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
290 |
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
291 |
text {*
|
| 99 | 292 |
To see what happens, assume the function defined in Figure~\ref{fig:sptac}, which
|
293 |
takes a record as argument and just prints out the content of this record (using the |
|
294 |
string transformation functions defined in Section~\ref{sec:printing}). Consider
|
|
295 |
now the proof |
|
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
296 |
*} |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
297 |
|
| 99 | 298 |
lemma shows "\<And>x y. A x y \<Longrightarrow> B y x \<longrightarrow> C (?z y) x" |
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
299 |
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
300 |
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
301 |
txt {*
|
| 99 | 302 |
which yields the printout: |
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
303 |
|
| 99 | 304 |
\begin{quote}\small
|
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
305 |
\begin{tabular}{ll}
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
306 |
params: & @{term x}, @{term y}\\
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
307 |
schematics: & @{term z}\\
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
308 |
asms: & @{term "A x y"}\\
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
309 |
concl: & @{term "B y x \<longrightarrow> C (z y) x"}\\
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
310 |
prems: & @{term "A x y"}
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
311 |
\end{tabular}
|
| 99 | 312 |
\end{quote}
|
313 |
||
314 |
Note in the actual output the brown colour of the variables @{term x} and
|
|
315 |
@{term y}. Although parameters in the original goal, they are fixed inside
|
|
316 |
the subproof. Similarly the schematic variable @{term z}. The assumption
|
|
317 |
is bound once as @{ML_type cterm} to the record-variable @{text asms} and
|
|
318 |
another time as @{ML_type thm} to @{text prems}.
|
|
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
319 |
|
| 99 | 320 |
Notice also that we had to append @{text "?"} to \isacommand{apply}. The
|
321 |
reason is that @{ML SUBPROOF} normally expects that the subgoal is solved completely.
|
|
322 |
Since in the function @{ML sp_tac} we returned the tactic @{ML no_tac}, the subproof
|
|
323 |
obviously fails. The question-mark allows us to recover from this failure |
|
324 |
in a graceful manner so that the warning messages are not overwritten |
|
325 |
by the error message. |
|
326 |
||
327 |
If we continue the proof script by applying the @{text impI}-rule
|
|
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
328 |
*} |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
329 |
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
330 |
apply(rule impI) |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
331 |
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
332 |
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
333 |
txt {*
|
| 99 | 334 |
then @{ML SUBPROOF} prints out
|
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
335 |
|
| 99 | 336 |
\begin{quote}\small
|
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
337 |
\begin{tabular}{ll}
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
338 |
params: & @{term x}, @{term y}\\
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
339 |
schematics: & @{term z}\\
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
340 |
asms: & @{term "A x y"}, @{term "B y x"}\\
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
341 |
concl: & @{term "C (z y) x"}\\
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
342 |
prems: & @{term "A x y"}, @{term "B y x"}
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
343 |
\end{tabular}
|
| 99 | 344 |
\end{quote}
|
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
345 |
*} |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
346 |
(*<*)oops(*>*) |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
347 |
|
| 99 | 348 |
text {*
|
349 |
where we now also have @{term "B y x"} as assumption.
|
|
350 |
||
351 |
One convenience of @{ML SUBPROOF} is that we can apply assumption
|
|
352 |
using the usual tactics, because the parameter @{text prems}
|
|
353 |
contains the assumptions as theorems. With this we can easily |
|
354 |
implement a tactic that almost behaves like @{ML atac}:
|
|
355 |
*} |
|
356 |
||
357 |
lemma shows "\<And>x y. \<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y" |
|
358 |
apply(tactic |
|
359 |
{* SUBPROOF (fn {prems, ...} => resolve_tac prems 1) @{context} 1 *})
|
|
360 |
txt{* yields
|
|
361 |
@{subgoals [display]} *}
|
|
362 |
(*<*)oops(*>*) |
|
363 |
||
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
364 |
|
| 93 | 365 |
text {*
|
| 99 | 366 |
The restriction in this tactic is that it cannot instantiate any |
367 |
schematic variables. This might be seen as a defect, but is actually |
|
368 |
an advantage in the situations for which @{ML SUBPROOF} was designed:
|
|
369 |
the reason is that instantiation of schematic variables can potentially |
|
370 |
has affect several goals and can render them unprovable. |
|
371 |
||
372 |
A similar but less powerful function is @{ML SUBGOAL}. It allows you to
|
|
373 |
inspect a subgoal specified by a number. |
|
| 93 | 374 |
*} |
375 |
||
| 99 | 376 |
ML %linenumbers{*fun select_tac (t,i) =
|
377 |
case t of |
|
378 |
@{term "Trueprop"} $ t' => select_tac (t',i)
|
|
379 |
| @{term "op \<and>"} $ _ $ _ => rtac @{thm conjI} i
|
|
380 |
| @{term "op \<longrightarrow>"} $ _ $ _ => rtac @{thm impI} i
|
|
381 |
| @{term "Not"} $ _ => rtac @{thm notI} i
|
|
382 |
| Const (@{const_name "All"}, _) $ _ => rtac @{thm allI} i
|
|
383 |
| _ => no_tac*} |
|
384 |
||
385 |
lemma shows "A \<and> B" "A \<longrightarrow> B" "\<forall>x. C x" |
|
386 |
apply(tactic {* SUBGOAL select_tac 1 *})
|
|
387 |
apply(tactic {* SUBGOAL select_tac 3 *})
|
|
388 |
apply(tactic {* SUBGOAL select_tac 4 *})
|
|
389 |
txt{* @{subgoals [display]} *}
|
|
390 |
(*<*)oops(*>*) |
|
391 |
||
392 |
text {*
|
|
393 |
However, this example is contrived, as there are much simpler ways |
|
394 |
to implement a proof procedure like the one above. They will be explained |
|
395 |
in the next section. |
|
396 |
||
397 |
A variant of @{ML SUBGOAL} is @{ML CSUBGOAL} which allows access to the goal
|
|
398 |
as @{ML_type cterm} instead of a @{ML_type term}.
|
|
| 93 | 399 |
*} |
400 |
||
| 99 | 401 |
|
| 93 | 402 |
section {* Operations on Tactics *}
|
403 |
||
| 99 | 404 |
text {* @{ML THEN} *}
|
| 93 | 405 |
|
| 99 | 406 |
lemma shows "(Foo \<and> Bar) \<and> False" |
407 |
apply(tactic {* (rtac @{thm conjI} 1)
|
|
408 |
THEN (rtac @{thm conjI} 1) *})
|
|
| 93 | 409 |
txt {* @{subgoals [display]} *}
|
410 |
(*<*)oops(*>*) |
|
411 |
||
| 99 | 412 |
ML{*val orelse_xmp = (rtac @{thm disjI1} ORELSE' rtac @{thm conjI})*}
|
413 |
||
414 |
lemma shows "True \<and> False" and "Foo \<or> Bar" |
|
415 |
apply(tactic {* orelse_xmp 1 *})
|
|
416 |
apply(tactic {* orelse_xmp 3 *})
|
|
| 93 | 417 |
txt {* @{subgoals [display]} *}
|
418 |
(*<*)oops(*>*) |
|
419 |
||
420 |
||
421 |
text {*
|
|
| 99 | 422 |
@{ML EVERY} @{ML REPEAT} @{ML DETERM}
|
| 93 | 423 |
|
424 |
@{ML rewrite_goals_tac}
|
|
425 |
@{ML cut_facts_tac}
|
|
426 |
@{ML ObjectLogic.full_atomize_tac}
|
|
427 |
@{ML ObjectLogic.rulify_tac}
|
|
428 |
@{ML resolve_tac}
|
|
429 |
*} |
|
430 |
||
431 |
||
432 |
||
433 |
text {*
|
|
434 |
||
435 |
||
436 |
A goal (or goal state) is a special @{ML_type thm}, which by
|
|
437 |
convention is an implication of the form: |
|
438 |
||
439 |
@{text[display] "A\<^isub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^isub>n \<Longrightarrow> #(C)"}
|
|
440 |
||
441 |
where @{term C} is the goal to be proved and the @{term "A\<^isub>i"} are the open
|
|
442 |
subgoals. |
|
443 |
Since the goal @{term C} can potentially be an implication, there is a
|
|
444 |
@{text "#"} wrapped around it, which prevents that premises are
|
|
445 |
misinterpreted as open subgoals. The wrapper @{text "# :: prop \<Rightarrow>
|
|
446 |
prop"} is just the identity function and used as a syntactic marker. |
|
447 |
||
448 |
||
449 |
||
450 |
||
451 |
||
452 |
While tactics can operate on the subgoals (the @{text "A\<^isub>i"} above), they
|
|
453 |
are expected to leave the conclusion @{term C} intact, with the
|
|
454 |
exception of possibly instantiating schematic variables. |
|
455 |
||
456 |
||
457 |
||
458 |
*} |
|
459 |
||
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
460 |
section {* Structured Proofs *}
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
461 |
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
462 |
lemma True |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
463 |
proof |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
464 |
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
465 |
{
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
466 |
fix A B C |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
467 |
assume r: "A & B \<Longrightarrow> C" |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
468 |
assume A B |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
469 |
then have "A & B" .. |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
470 |
then have C by (rule r) |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
471 |
} |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
472 |
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
473 |
{
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
474 |
fix A B C |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
475 |
assume r: "A & B \<Longrightarrow> C" |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
476 |
assume A B |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
477 |
note conjI [OF this] |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
478 |
note r [OF this] |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
479 |
} |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
480 |
oops |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
481 |
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
482 |
ML {* fun prop ctxt s =
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
483 |
Thm.cterm_of (ProofContext.theory_of ctxt) (Syntax.read_prop ctxt s) *} |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
484 |
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
485 |
ML {*
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
486 |
val ctxt0 = @{context};
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
487 |
val ctxt = ctxt0; |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
488 |
val (_, ctxt) = Variable.add_fixes ["A", "B", "C"] ctxt; |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
489 |
val ([r], ctxt) = Assumption.add_assumes [prop ctxt "A & B \<Longrightarrow> C"] ctxt; |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
490 |
val (this, ctxt) = Assumption.add_assumes [prop ctxt "A", prop ctxt "B"] ctxt; |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
491 |
val this = [@{thm conjI} OF this];
|
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
492 |
val this = r OF this; |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
493 |
val this = Assumption.export false ctxt ctxt0 this |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
494 |
val this = Variable.export ctxt ctxt0 [this] |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
495 |
*} |
| 93 | 496 |
|
497 |
||
498 |
end |