author | boehmes |
Wed, 25 Feb 2009 10:14:42 +0100 | |
changeset 141 | 5aa3140ad52e |
parent 139 | ed1eb9cb2533 |
parent 137 | a9685909944d |
child 142 | c06885c36575 |
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 {* |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
8 |
The main reason for descending to the ML-level of Isabelle is to be able to |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
9 |
implement automatic proof procedures. Such proof procedures usually lessen |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
10 |
considerably the burden of manual reasoning, for example, when introducing |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
11 |
new definitions. These proof procedures are centred around refining a goal |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
12 |
state using tactics. This is similar to the \isacommand{apply}-style |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
13 |
reasoning at the user level, where goals are modified in a sequence of proof |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
14 |
steps until all of them are solved. However, there are also more structured |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
15 |
operations available on the ML-level that help with the handling of |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
16 |
variables and assumptions. |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
17 |
|
93 | 18 |
*} |
19 |
||
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
20 |
section {* Basics of Reasoning with Tactics*} |
93 | 21 |
|
22 |
text {* |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
23 |
To see how tactics work, let us first transcribe a simple \isacommand{apply}-style proof |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
24 |
into ML. Suppose the following proof. |
93 | 25 |
*} |
26 |
||
27 |
lemma disj_swap: "P \<or> Q \<Longrightarrow> Q \<or> P" |
|
28 |
apply(erule disjE) |
|
29 |
apply(rule disjI2) |
|
30 |
apply(assumption) |
|
31 |
apply(rule disjI1) |
|
32 |
apply(assumption) |
|
33 |
done |
|
34 |
||
35 |
text {* |
|
36 |
This proof translates to the following ML-code. |
|
37 |
||
38 |
@{ML_response_fake [display,gray] |
|
39 |
"let |
|
40 |
val ctxt = @{context} |
|
41 |
val goal = @{prop \"P \<or> Q \<Longrightarrow> Q \<or> P\"} |
|
42 |
in |
|
43 |
Goal.prove ctxt [\"P\", \"Q\"] [] goal |
|
44 |
(fn _ => |
|
45 |
etac @{thm disjE} 1 |
|
46 |
THEN rtac @{thm disjI2} 1 |
|
47 |
THEN atac 1 |
|
48 |
THEN rtac @{thm disjI1} 1 |
|
49 |
THEN atac 1) |
|
50 |
end" "?P \<or> ?Q \<Longrightarrow> ?Q \<or> ?P"} |
|
51 |
||
52 |
To start the proof, the function @{ML "Goal.prove"}~@{text "ctxt xs As C |
|
99 | 53 |
tac"} sets up a goal state for proving the goal @{text C} |
54 |
(that is @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"} in the proof at hand) under the |
|
55 |
assumptions @{text As} (happens to be empty) with the variables |
|
93 | 56 |
@{text xs} that will be generalised once the goal is proved (in our case |
57 |
@{text P} and @{text Q}). The @{text "tac"} is the tactic that proves the goal; |
|
58 |
it can make use of the local assumptions (there are none in this example). |
|
120
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
59 |
The functions @{ML etac}, @{ML rtac} and @{ML atac} in the code above correspond to |
93 | 60 |
@{text erule}, @{text rule} and @{text assumption}, respectively. |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
61 |
The operator @{ML THEN} strings the tactics together. |
93 | 62 |
|
63 |
\begin{readmore} |
|
99 | 64 |
To learn more about the function @{ML Goal.prove} see \isccite{sec:results} |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
65 |
and the file @{ML_file "Pure/goal.ML"}. See @{ML_file |
99 | 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 |
||
120
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
71 |
Note that in the code above we use antiquotations for referencing the theorems. Many theorems |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
72 |
also have ML-bindings with the same name. Therefore, we could also just have |
118
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
73 |
written @{ML "etac disjE 1"}, or in case where there are no ML-binding obtain |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
74 |
the theorem dynamically using the function @{ML thm}; for example |
109 | 75 |
\mbox{@{ML "etac (thm \"disjE\") 1"}}. Both ways however are considered bad style! |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
76 |
The reason |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
77 |
is that the binding for @{ML disjE} can be re-assigned by the user and thus |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
78 |
one does not have complete control over which theorem is actually |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
79 |
applied. This problem is nicely prevented by using antiquotations, because |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
80 |
then the theorems are fixed statically at compile-time. |
93 | 81 |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
82 |
During the development of automatic proof procedures, you will often find it |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
83 |
necessary to test a tactic on examples. This can be conveniently |
93 | 84 |
done with the command \isacommand{apply}@{text "(tactic \<verbopen> \<dots> \<verbclose>)"}. |
85 |
Consider the following sequence of tactics |
|
86 |
*} |
|
87 |
||
88 |
ML{*val foo_tac = |
|
89 |
(etac @{thm disjE} 1 |
|
90 |
THEN rtac @{thm disjI2} 1 |
|
91 |
THEN atac 1 |
|
92 |
THEN rtac @{thm disjI1} 1 |
|
93 |
THEN atac 1)*} |
|
94 |
||
95 |
text {* and the Isabelle proof: *} |
|
96 |
||
97 |
lemma "P \<or> Q \<Longrightarrow> Q \<or> P" |
|
98 |
apply(tactic {* foo_tac *}) |
|
99 |
done |
|
100 |
||
101 |
text {* |
|
104 | 102 |
By using @{text "tactic \<verbopen> \<dots> \<verbclose>"} you can call from the |
103 |
user level of Isabelle the tactic @{ML foo_tac} or |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
104 |
any other function that returns a tactic. |
93 | 105 |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
106 |
The tactic @{ML foo_tac} is just a sequence of simple tactics stringed |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
107 |
together by @{ML THEN}. As can be seen, each simple tactic in @{ML foo_tac} |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
108 |
has a hard-coded number that stands for the subgoal analysed by the |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
109 |
tactic (@{text "1"} stands for the first, or top-most, subgoal). This hard-coding |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
110 |
of goals is sometimes wanted, but usually it is not. To avoid the explicit numbering, |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
111 |
you can write\label{tac:footacprime} |
93 | 112 |
*} |
113 |
||
114 |
ML{*val foo_tac' = |
|
115 |
(etac @{thm disjE} |
|
116 |
THEN' rtac @{thm disjI2} |
|
117 |
THEN' atac |
|
118 |
THEN' rtac @{thm disjI1} |
|
119 |
THEN' atac)*} |
|
120 |
||
121 |
text {* |
|
122 |
and then give the number for the subgoal explicitly when the tactic is |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
123 |
called. So in the next proof you can first discharge the second subgoal, and |
109 | 124 |
subsequently the first. |
93 | 125 |
*} |
126 |
||
127 |
lemma "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1" |
|
128 |
and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2" |
|
129 |
apply(tactic {* foo_tac' 2 *}) |
|
130 |
apply(tactic {* foo_tac' 1 *}) |
|
131 |
done |
|
132 |
||
133 |
text {* |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
134 |
This kind of addressing is more difficult to achieve when the goal is |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
135 |
hard-coded inside the tactic. For most operators that combine tactics |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
136 |
(@{ML THEN} is only one such operator) a ``primed'' version exists. |
99 | 137 |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
138 |
The tactics @{ML foo_tac} and @{ML foo_tac'} are very specific for |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
139 |
analysing goals being only of the form @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. If the goal is not |
109 | 140 |
of this form, then they return the error message: |
99 | 141 |
|
142 |
\begin{isabelle} |
|
143 |
@{text "*** empty result sequence -- proof command failed"}\\ |
|
144 |
@{text "*** At command \"apply\"."} |
|
145 |
\end{isabelle} |
|
146 |
||
109 | 147 |
This means the tactics failed. The reason for this error message is that tactics |
148 |
are functions mapping a goal state to a (lazy) sequence of successor states. |
|
149 |
Hence the type of a tactic is: |
|
150 |
*} |
|
93 | 151 |
|
109 | 152 |
ML{*type tactic = thm -> thm Seq.seq*} |
93 | 153 |
|
109 | 154 |
text {* |
155 |
By convention, if a tactic fails, then it should return the empty sequence. |
|
156 |
Therefore, if you write your own tactics, they should not raise exceptions |
|
157 |
willy-nilly; only in very grave failure situations should a tactic raise the |
|
158 |
exception @{text THM}. |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
159 |
|
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
160 |
The simplest tactics are @{ML no_tac} and @{ML all_tac}. The first returns |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
161 |
the empty sequence and is defined as |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
162 |
*} |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
163 |
|
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
164 |
ML{*fun no_tac thm = Seq.empty*} |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
165 |
|
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
166 |
text {* |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
167 |
which means @{ML no_tac} always fails. The second returns the given theorem wrapped |
109 | 168 |
up in a single member sequence; it is defined as |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
169 |
*} |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
170 |
|
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
171 |
ML{*fun all_tac thm = Seq.single thm*} |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
172 |
|
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
173 |
text {* |
109 | 174 |
which means @{ML all_tac} always succeeds, but also does not make any progress |
175 |
with the proof. |
|
93 | 176 |
|
109 | 177 |
The lazy list of possible successor goal states shows through at the user-level |
99 | 178 |
of Isabelle when using the command \isacommand{back}. For instance in the |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
179 |
following proof there are two possibilities for how to apply |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
180 |
@{ML foo_tac'}: either using the first assumption or the second. |
93 | 181 |
*} |
182 |
||
183 |
lemma "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P" |
|
184 |
apply(tactic {* foo_tac' 1 *}) |
|
185 |
back |
|
186 |
done |
|
187 |
||
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
188 |
|
93 | 189 |
text {* |
99 | 190 |
By using \isacommand{back}, we construct the proof that uses the |
109 | 191 |
second assumption. While in the proof above, it does not really matter which |
192 |
assumption is used, in more interesting cases provability might depend |
|
193 |
on exploring different possibilities. |
|
99 | 194 |
|
93 | 195 |
\begin{readmore} |
196 |
See @{ML_file "Pure/General/seq.ML"} for the implementation of lazy |
|
109 | 197 |
sequences. In day-to-day Isabelle programming, however, one rarely |
198 |
constructs sequences explicitly, but uses the predefined tactics and |
|
199 |
tactic combinators instead. |
|
93 | 200 |
\end{readmore} |
201 |
||
104 | 202 |
It might be surprising that tactics, which transform |
109 | 203 |
one goal state to the next, are functions from theorems to theorem |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
204 |
(sequences). The surprise resolves by knowing that every |
104 | 205 |
goal state is indeed a theorem. To shed more light on this, |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
206 |
let us modify the code of @{ML all_tac} to obtain the following |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
207 |
tactic |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
208 |
*} |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
209 |
|
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
210 |
ML{*fun my_print_tac ctxt thm = |
132 | 211 |
let |
212 |
val _ = warning (str_of_thm ctxt thm) |
|
213 |
in |
|
214 |
Seq.single thm |
|
215 |
end*} |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
216 |
|
109 | 217 |
text_raw {* |
218 |
\begin{figure}[p] |
|
219 |
\begin{isabelle} |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
220 |
*} |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
221 |
lemma shows "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B" |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
222 |
apply(tactic {* my_print_tac @{context} *}) |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
223 |
|
109 | 224 |
txt{* \begin{minipage}{\textwidth} |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
225 |
@{subgoals [display]} |
109 | 226 |
\end{minipage}\medskip |
227 |
||
228 |
\begin{minipage}{\textwidth} |
|
229 |
\small\colorbox{gray!20}{ |
|
230 |
\begin{tabular}{@ {}l@ {}} |
|
231 |
internal goal state:\\ |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
232 |
@{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"} |
109 | 233 |
\end{tabular}} |
234 |
\end{minipage}\medskip |
|
93 | 235 |
*} |
236 |
||
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
237 |
apply(rule conjI) |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
238 |
apply(tactic {* my_print_tac @{context} *}) |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
239 |
|
109 | 240 |
txt{* \begin{minipage}{\textwidth} |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
241 |
@{subgoals [display]} |
109 | 242 |
\end{minipage}\medskip |
243 |
||
244 |
\begin{minipage}{\textwidth} |
|
245 |
\small\colorbox{gray!20}{ |
|
246 |
\begin{tabular}{@ {}l@ {}} |
|
247 |
internal goal state:\\ |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
248 |
@{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"} |
109 | 249 |
\end{tabular}} |
250 |
\end{minipage}\medskip |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
251 |
*} |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
252 |
|
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
253 |
apply(assumption) |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
254 |
apply(tactic {* my_print_tac @{context} *}) |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
255 |
|
109 | 256 |
txt{* \begin{minipage}{\textwidth} |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
257 |
@{subgoals [display]} |
109 | 258 |
\end{minipage}\medskip |
259 |
||
260 |
\begin{minipage}{\textwidth} |
|
261 |
\small\colorbox{gray!20}{ |
|
262 |
\begin{tabular}{@ {}l@ {}} |
|
263 |
internal goal state:\\ |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
264 |
@{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"} |
109 | 265 |
\end{tabular}} |
266 |
\end{minipage}\medskip |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
267 |
*} |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
268 |
|
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
269 |
apply(assumption) |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
270 |
apply(tactic {* my_print_tac @{context} *}) |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
271 |
|
109 | 272 |
txt{* \begin{minipage}{\textwidth} |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
273 |
@{subgoals [display]} |
109 | 274 |
\end{minipage}\medskip |
275 |
||
276 |
\begin{minipage}{\textwidth} |
|
277 |
\small\colorbox{gray!20}{ |
|
278 |
\begin{tabular}{@ {}l@ {}} |
|
279 |
internal goal state:\\ |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
280 |
@{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"} |
109 | 281 |
\end{tabular}} |
282 |
\end{minipage}\medskip |
|
283 |
*} |
|
284 |
done |
|
285 |
text_raw {* |
|
286 |
\end{isabelle} |
|
118
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
287 |
\caption{The figure shows a proof where each intermediate goal state is |
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
288 |
printed by the Isabelle system and by @{ML my_print_tac}. The latter shows |
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
289 |
the goal state as represented internally (highlighted boxes). This |
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
290 |
illustrates that every goal state in Isabelle is represented by a theorem: |
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
291 |
when we start the proof of \mbox{@{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}} the theorem is |
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
292 |
@{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}; when we finish the proof the |
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
293 |
theorem is @{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}.\label{fig:goalstates}} |
109 | 294 |
\end{figure} |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
295 |
*} |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
296 |
|
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
297 |
|
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
298 |
text {* |
109 | 299 |
which prints out the given theorem (using the string-function defined in |
300 |
Section~\ref{sec:printing}) and then behaves like @{ML all_tac}. With this |
|
301 |
tactic we are in the position to inspect every goal state in a |
|
302 |
proof. Consider now the proof in Figure~\ref{fig:goalstates}: as can be seen, |
|
303 |
internally every goal state is an implication of the form |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
304 |
|
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
305 |
@{text[display] "A\<^isub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^isub>n \<Longrightarrow> (C)"} |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
306 |
|
109 | 307 |
where @{term C} is the goal to be proved and the @{term "A\<^isub>i"} are |
308 |
the subgoals. So after setting up the lemma, the goal state is always of the |
|
309 |
form @{text "C \<Longrightarrow> (C)"}; when the proof is finished we are left with @{text |
|
310 |
"(C)"}. Since the goal @{term C} can potentially be an implication, there is |
|
311 |
a ``protector'' wrapped around it (in from of an outermost constant @{text |
|
120
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
312 |
"Const (\"prop\", bool \<Rightarrow> bool)"}; however this constant |
109 | 313 |
is invisible in the figure). This prevents that premises of @{text C} are |
314 |
mis-interpreted as open subgoals. While tactics can operate on the subgoals |
|
315 |
(the @{text "A\<^isub>i"} above), they are expected to leave the conclusion |
|
316 |
@{term C} intact, with the exception of possibly instantiating schematic |
|
317 |
variables. If you use the predefined tactics, which we describe in the next |
|
318 |
section, this will always be the case. |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
319 |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
320 |
\begin{readmore} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
321 |
For more information about the internals of goals see \isccite{sec:tactical-goals}. |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
322 |
\end{readmore} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
323 |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
324 |
*} |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
325 |
|
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
326 |
section {* Simple Tactics *} |
93 | 327 |
|
99 | 328 |
text {* |
109 | 329 |
Let us start with the tactic @{ML print_tac}, which is quite useful for low-level |
330 |
debugging of tactics. It just prints out a message and the current goal state. |
|
331 |
Processing the proof |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
332 |
*} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
333 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
334 |
lemma shows "False \<Longrightarrow> True" |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
335 |
apply(tactic {* print_tac "foo message" *}) |
109 | 336 |
txt{*gives:\medskip |
337 |
||
338 |
\begin{minipage}{\textwidth}\small |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
339 |
@{text "foo message"}\\[3mm] |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
340 |
@{prop "False \<Longrightarrow> True"}\\ |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
341 |
@{text " 1. False \<Longrightarrow> True"}\\ |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
342 |
\end{minipage} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
343 |
*} |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
344 |
(*<*)oops(*>*) |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
345 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
346 |
text {* |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
347 |
Another simple tactic is the function @{ML atac}, which, as shown in the previous |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
348 |
section, corresponds to the assumption command. |
99 | 349 |
*} |
350 |
||
351 |
lemma shows "P \<Longrightarrow> P" |
|
93 | 352 |
apply(tactic {* atac 1 *}) |
109 | 353 |
txt{*\begin{minipage}{\textwidth} |
354 |
@{subgoals [display]} |
|
355 |
\end{minipage}*} |
|
356 |
(*<*)oops(*>*) |
|
93 | 357 |
|
99 | 358 |
text {* |
359 |
Similarly, @{ML rtac}, @{ML dtac}, @{ML etac} and @{ML ftac} correspond |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
360 |
to @{text rule}, @{text drule}, @{text erule} and @{text frule}, |
109 | 361 |
respectively. Each of them takes a theorem as argument and attempts to |
362 |
apply it to a goal. Below are three self-explanatory examples. |
|
99 | 363 |
*} |
364 |
||
365 |
lemma shows "P \<and> Q" |
|
93 | 366 |
apply(tactic {* rtac @{thm conjI} 1 *}) |
104 | 367 |
txt{*\begin{minipage}{\textwidth} |
368 |
@{subgoals [display]} |
|
369 |
\end{minipage}*} |
|
93 | 370 |
(*<*)oops(*>*) |
371 |
||
99 | 372 |
lemma shows "P \<and> Q \<Longrightarrow> False" |
93 | 373 |
apply(tactic {* etac @{thm conjE} 1 *}) |
104 | 374 |
txt{*\begin{minipage}{\textwidth} |
375 |
@{subgoals [display]} |
|
376 |
\end{minipage}*} |
|
93 | 377 |
(*<*)oops(*>*) |
378 |
||
379 |
lemma shows "False \<and> True \<Longrightarrow> False" |
|
380 |
apply(tactic {* dtac @{thm conjunct2} 1 *}) |
|
104 | 381 |
txt{*\begin{minipage}{\textwidth} |
382 |
@{subgoals [display]} |
|
383 |
\end{minipage}*} |
|
93 | 384 |
(*<*)oops(*>*) |
385 |
||
386 |
text {* |
|
109 | 387 |
Note the number in each tactic call. Also as mentioned in the previous section, most |
388 |
basic tactics take such an argument; it addresses the subgoal they are analysing. |
|
389 |
In the proof below, we first split up the conjunction in the second subgoal by |
|
390 |
focusing on this subgoal first. |
|
99 | 391 |
*} |
392 |
||
393 |
lemma shows "Foo" and "P \<and> Q" |
|
394 |
apply(tactic {* rtac @{thm conjI} 2 *}) |
|
104 | 395 |
txt {*\begin{minipage}{\textwidth} |
396 |
@{subgoals [display]} |
|
397 |
\end{minipage}*} |
|
99 | 398 |
(*<*)oops(*>*) |
399 |
||
109 | 400 |
text {* |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
401 |
The function @{ML resolve_tac} is similar to @{ML rtac}, except that it |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
402 |
expects a list of theorems as arguments. From this list it will apply the |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
403 |
first applicable theorem (later theorems that are also applicable can be |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
404 |
explored via the lazy sequences mechanism). Given the code |
93 | 405 |
*} |
406 |
||
99 | 407 |
ML{*val resolve_tac_xmp = resolve_tac [@{thm impI}, @{thm conjI}]*} |
408 |
||
409 |
text {* |
|
410 |
an example for @{ML resolve_tac} is the following proof where first an outermost |
|
411 |
implication is analysed and then an outermost conjunction. |
|
412 |
*} |
|
413 |
||
414 |
lemma shows "C \<longrightarrow> (A \<and> B)" and "(A \<longrightarrow> B) \<and> C" |
|
415 |
apply(tactic {* resolve_tac_xmp 1 *}) |
|
416 |
apply(tactic {* resolve_tac_xmp 2 *}) |
|
104 | 417 |
txt{*\begin{minipage}{\textwidth} |
418 |
@{subgoals [display]} |
|
419 |
\end{minipage}*} |
|
99 | 420 |
(*<*)oops(*>*) |
421 |
||
422 |
text {* |
|
109 | 423 |
Similarly versions taking a list of theorems exist for the tactics |
424 |
@{ML dtac} (@{ML dresolve_tac}), @{ML etac} (@{ML eresolve_tac}) and so on. |
|
425 |
||
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
426 |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
427 |
Another simple tactic is @{ML cut_facts_tac}. It inserts a list of theorems |
109 | 428 |
into the assumptions of the current goal state. For example |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
429 |
*} |
99 | 430 |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
431 |
lemma shows "True \<noteq> False" |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
432 |
apply(tactic {* cut_facts_tac [@{thm True_def}, @{thm False_def}] 1 *}) |
109 | 433 |
txt{*produces the goal state\medskip |
434 |
||
435 |
\begin{minipage}{\textwidth} |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
436 |
@{subgoals [display]} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
437 |
\end{minipage}*} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
438 |
(*<*)oops(*>*) |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
439 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
440 |
text {* |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
441 |
Since rules are applied using higher-order unification, an automatic proof |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
442 |
procedure might become too fragile, if it just applies inference rules as |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
443 |
shown above. The reason is that a number of rules introduce meta-variables |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
444 |
into the goal state. Consider for example the proof |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
445 |
*} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
446 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
447 |
lemma shows "\<forall>x\<in>A. P x \<Longrightarrow> Q x" |
118
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
448 |
apply(tactic {* dtac @{thm bspec} 1 *}) |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
449 |
txt{*\begin{minipage}{\textwidth} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
450 |
@{subgoals [display]} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
451 |
\end{minipage}*} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
452 |
(*<*)oops(*>*) |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
453 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
454 |
text {* |
109 | 455 |
where the application of Rule @{text bspec} generates two subgoals involving the |
456 |
meta-variable @{text "?x"}. Now, if you are not careful, tactics |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
457 |
applied to the first subgoal might instantiate this meta-variable in such a |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
458 |
way that the second subgoal becomes unprovable. If it is clear what the @{text "?x"} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
459 |
should be, then this situation can be avoided by introducing a more |
109 | 460 |
constraint version of the @{text bspec}-rule. Such constraints can be given by |
461 |
pre-instantiating theorems with other theorems. One function to do this is |
|
462 |
@{ML RS} |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
463 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
464 |
@{ML_response_fake [display,gray] |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
465 |
"@{thm disjI1} RS @{thm conjI}" "\<lbrakk>?P1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<or> ?Q1) \<and> ?Q"} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
466 |
|
109 | 467 |
which in the example instantiates the first premise of the @{text conjI}-rule |
468 |
with the rule @{text disjI1}. If the instantiation is impossible, as in the |
|
469 |
case of |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
470 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
471 |
@{ML_response_fake_both [display,gray] |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
472 |
"@{thm conjI} RS @{thm mp}" |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
473 |
"*** Exception- THM (\"RSN: no unifiers\", 1, |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
474 |
[\"\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q\", \"\<lbrakk>?P \<longrightarrow> ?Q; ?P\<rbrakk> \<Longrightarrow> ?Q\"]) raised"} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
475 |
|
109 | 476 |
then the function raises an exception. The function @{ML RSN} is similar to @{ML RS}, but |
477 |
takes an additional number as argument that makes explicit which premise |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
478 |
should be instantiated. |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
479 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
480 |
To improve readability of the theorems we produce below, we shall use |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
481 |
the following function |
99 | 482 |
*} |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
483 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
484 |
ML{*fun no_vars ctxt thm = |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
485 |
let |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
486 |
val ((_, [thm']), _) = Variable.import_thms true [thm] ctxt |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
487 |
in |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
488 |
thm' |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
489 |
end*} |
93 | 490 |
|
99 | 491 |
text {* |
109 | 492 |
that transform the schematic variables of a theorem into free variables. |
493 |
Using this function for the first @{ML RS}-expression above would produce |
|
494 |
the more readable result: |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
495 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
496 |
@{ML_response_fake [display,gray] |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
497 |
"no_vars @{context} (@{thm disjI1} RS @{thm conjI})" "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> (P \<or> Qa) \<and> Q"} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
498 |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
499 |
If you want to instantiate more than one premise of a theorem, you can use |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
500 |
the function @{ML MRS}: |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
501 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
502 |
@{ML_response_fake [display,gray] |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
503 |
"no_vars @{context} ([@{thm disjI1}, @{thm disjI2}] MRS @{thm conjI})" |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
504 |
"\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> (P \<or> Qa) \<and> (Pa \<or> Q)"} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
505 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
506 |
If you need to instantiate lists of theorems, you can use the |
109 | 507 |
functions @{ML RL} and @{ML MRL}. For example in the code below, |
508 |
every theorem in the second list is instantiated with every |
|
509 |
theorem in the first. |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
510 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
511 |
@{ML_response_fake [display,gray] |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
512 |
"[@{thm impI}, @{thm disjI2}] RL [@{thm conjI}, @{thm disjI1}]" |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
513 |
"[\<lbrakk>P \<Longrightarrow> Q; Qa\<rbrakk> \<Longrightarrow> (P \<longrightarrow> Q) \<and> Qa, |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
514 |
\<lbrakk>Q; Qa\<rbrakk> \<Longrightarrow> (P \<or> Q) \<and> Qa, |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
515 |
(P \<Longrightarrow> Q) \<Longrightarrow> (P \<longrightarrow> Q) \<or> Qa, |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
516 |
Q \<Longrightarrow> (P \<or> Q) \<or> Qa]"} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
517 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
518 |
\begin{readmore} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
519 |
The combinators for instantiating theorems are defined in @{ML_file "Pure/drule.ML"}. |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
520 |
\end{readmore} |
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
521 |
|
109 | 522 |
Often proofs on the ML-level involve elaborate operations on assumptions and |
523 |
@{text "\<And>"}-quantified variables. To do such operations using the basic tactics |
|
128 | 524 |
shown so far is very unwieldy and brittle. Some convenience and |
99 | 525 |
safety is provided by the tactic @{ML SUBPROOF}. This tactic fixes the parameters |
109 | 526 |
and binds the various components of a goal state to a record. |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
527 |
To see what happens, assume the function defined in Figure~\ref{fig:sptac}, which |
109 | 528 |
takes a record and just prints out the content of this record (using the |
529 |
string transformation functions from in Section~\ref{sec:printing}). Consider |
|
530 |
now the proof: |
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
531 |
*} |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
532 |
|
99 | 533 |
text_raw{* |
534 |
\begin{figure} |
|
535 |
\begin{isabelle} |
|
536 |
*} |
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
537 |
ML{*fun sp_tac {prems, params, asms, concl, context, schematics} = |
132 | 538 |
let |
539 |
val str_of_params = str_of_cterms context params |
|
540 |
val str_of_asms = str_of_cterms context asms |
|
541 |
val str_of_concl = str_of_cterm context concl |
|
542 |
val str_of_prems = str_of_thms context prems |
|
543 |
val str_of_schms = str_of_cterms context (snd schematics) |
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
544 |
|
132 | 545 |
val _ = (warning ("params: " ^ str_of_params); |
546 |
warning ("schematics: " ^ str_of_schms); |
|
547 |
warning ("assumptions: " ^ str_of_asms); |
|
548 |
warning ("conclusion: " ^ str_of_concl); |
|
549 |
warning ("premises: " ^ str_of_prems)) |
|
550 |
in |
|
551 |
no_tac |
|
552 |
end*} |
|
99 | 553 |
text_raw{* |
554 |
\end{isabelle} |
|
555 |
\caption{A function that prints out the various parameters provided by the tactic |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
556 |
@{ML SUBPROOF}. It uses the functions defined in Section~\ref{sec:printing} for |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
557 |
extracting strings from @{ML_type cterm}s and @{ML_type thm}s.\label{fig:sptac}} |
99 | 558 |
\end{figure} |
559 |
*} |
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
560 |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
561 |
|
99 | 562 |
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
|
563 |
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})? |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
564 |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
565 |
txt {* |
109 | 566 |
The tactic produces the following printout: |
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
567 |
|
99 | 568 |
\begin{quote}\small |
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
569 |
\begin{tabular}{ll} |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
570 |
params: & @{term x}, @{term y}\\ |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
571 |
schematics: & @{term z}\\ |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
572 |
assumptions: & @{term "A x y"}\\ |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
573 |
conclusion: & @{term "B y x \<longrightarrow> C (z y) x"}\\ |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
574 |
premises: & @{term "A x y"} |
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
575 |
\end{tabular} |
99 | 576 |
\end{quote} |
577 |
||
578 |
Note in the actual output the brown colour of the variables @{term x} and |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
579 |
@{term y}. Although they are parameters in the original goal, they are fixed inside |
109 | 580 |
the subproof. By convention these fixed variables are printed in brown colour. |
581 |
Similarly the schematic variable @{term z}. The assumption, or premise, |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
582 |
@{prop "A x y"} is bound as @{ML_type cterm} to the record-variable |
109 | 583 |
@{text asms}, but also 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
|
584 |
|
109 | 585 |
Notice also that we had to append @{text [quotes] "?"} to the |
586 |
\isacommand{apply}-command. The reason is that @{ML SUBPROOF} normally |
|
587 |
expects that the subgoal is solved completely. Since in the function @{ML |
|
588 |
sp_tac} we returned the tactic @{ML no_tac}, the subproof obviously |
|
589 |
fails. The question-mark allows us to recover from this failure in a |
|
590 |
graceful manner so that the warning messages are not overwritten by an |
|
591 |
``empty sequence'' error message. |
|
99 | 592 |
|
593 |
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
|
594 |
*} |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
595 |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
596 |
apply(rule impI) |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
597 |
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})? |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
598 |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
599 |
txt {* |
118
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
600 |
then the tactic prints out: |
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
601 |
|
99 | 602 |
\begin{quote}\small |
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
603 |
\begin{tabular}{ll} |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
604 |
params: & @{term x}, @{term y}\\ |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
605 |
schematics: & @{term z}\\ |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
606 |
assumptions: & @{term "A x y"}, @{term "B y x"}\\ |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
607 |
conclusion: & @{term "C (z y) x"}\\ |
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
608 |
premises: & @{term "A x y"}, @{term "B y x"} |
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
609 |
\end{tabular} |
99 | 610 |
\end{quote} |
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
611 |
*} |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
612 |
(*<*)oops(*>*) |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
613 |
|
99 | 614 |
text {* |
109 | 615 |
Now also @{term "B y x"} is an assumption bound to @{text asms} and @{text prems}. |
99 | 616 |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
617 |
One convenience of @{ML SUBPROOF} is that we can apply the assumptions |
99 | 618 |
using the usual tactics, because the parameter @{text prems} |
109 | 619 |
contains them as theorems. With this you can easily |
620 |
implement a tactic that behaves almost like @{ML atac}: |
|
99 | 621 |
*} |
622 |
||
104 | 623 |
ML{*val atac' = SUBPROOF (fn {prems, ...} => resolve_tac prems 1)*} |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
624 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
625 |
text {* |
109 | 626 |
If you apply @{ML atac'} to the next lemma |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
627 |
*} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
628 |
|
109 | 629 |
lemma shows "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y" |
104 | 630 |
apply(tactic {* atac' @{context} 1 *}) |
109 | 631 |
txt{* it will produce |
99 | 632 |
@{subgoals [display]} *} |
633 |
(*<*)oops(*>*) |
|
634 |
||
104 | 635 |
text {* |
109 | 636 |
The restriction in this tactic which is not present in @{ML atac} is |
637 |
that it cannot instantiate any |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
638 |
schematic variable. This might be seen as a defect, but it is actually |
104 | 639 |
an advantage in the situations for which @{ML SUBPROOF} was designed: |
109 | 640 |
the reason is that, as mentioned before, instantiation of schematic variables can affect |
104 | 641 |
several goals and can render them unprovable. @{ML SUBPROOF} is meant |
642 |
to avoid this. |
|
643 |
||
109 | 644 |
Notice that @{ML atac'} inside @{ML SUBPROOF} calls @{ML resolve_tac} with |
645 |
the subgoal number @{text "1"} and also the outer call to @{ML SUBPROOF} in |
|
646 |
the \isacommand{apply}-step uses @{text "1"}. This is another advantage |
|
647 |
of @{ML SUBPROOF}: the addressing inside it is completely |
|
648 |
local to the tactic inside the subproof. It is therefore possible to also apply |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
649 |
@{ML atac'} to the second goal by just writing: |
104 | 650 |
*} |
651 |
||
109 | 652 |
lemma shows "True" and "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y" |
104 | 653 |
apply(tactic {* atac' @{context} 2 *}) |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
654 |
apply(rule TrueI) |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
655 |
done |
104 | 656 |
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
657 |
|
93 | 658 |
text {* |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
659 |
\begin{readmore} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
660 |
The function @{ML SUBPROOF} is defined in @{ML_file "Pure/subgoal.ML"} and |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
661 |
also described in \isccite{sec:results}. |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
662 |
\end{readmore} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
663 |
|
104 | 664 |
A similar but less powerful function than @{ML SUBPROOF} is @{ML SUBGOAL}. |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
665 |
It allows you to inspect a given subgoal. With this you can implement |
109 | 666 |
a tactic that applies a rule according to the topmost logic connective in the |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
667 |
subgoal (to illustrate this we only analyse a few connectives). The code |
109 | 668 |
of the tactic is as follows.\label{tac:selecttac} |
93 | 669 |
*} |
670 |
||
114
13fd0a83d3c3
properly handled linenumbers in ML-text and Isar-proofs
Christian Urban <urbanc@in.tum.de>
parents:
109
diff
changeset
|
671 |
ML %linenosgray{*fun select_tac (t,i) = |
99 | 672 |
case t of |
673 |
@{term "Trueprop"} $ t' => select_tac (t',i) |
|
109 | 674 |
| @{term "op \<Longrightarrow>"} $ _ $ t' => select_tac (t',i) |
99 | 675 |
| @{term "op \<and>"} $ _ $ _ => rtac @{thm conjI} i |
676 |
| @{term "op \<longrightarrow>"} $ _ $ _ => rtac @{thm impI} i |
|
677 |
| @{term "Not"} $ _ => rtac @{thm notI} i |
|
678 |
| Const (@{const_name "All"}, _) $ _ => rtac @{thm allI} i |
|
104 | 679 |
| _ => all_tac*} |
99 | 680 |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
681 |
text {* |
109 | 682 |
The input of the function is a term representing the subgoal and a number |
683 |
specifying the subgoal of interest. In line 3 you need to descend under the |
|
684 |
outermost @{term "Trueprop"} in order to get to the connective you like to |
|
685 |
analyse. Otherwise goals like @{prop "A \<and> B"} are not properly |
|
686 |
analysed. Similarly with meta-implications in the next line. While for the |
|
687 |
first five patterns we can use the @{text "@term"}-antiquotation to |
|
688 |
construct the patterns, the pattern in Line 8 cannot be constructed in this |
|
689 |
way. The reason is that an antiquotation would fix the type of the |
|
690 |
quantified variable. So you really have to construct the pattern using the |
|
691 |
basic term-constructors. This is not necessary in other cases, because their type |
|
118
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
692 |
is always fixed to function types involving only the type @{typ bool}. For the |
109 | 693 |
final pattern, we chose to just return @{ML all_tac}. Consequently, |
694 |
@{ML select_tac} never fails. |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
695 |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
696 |
Let us now see how to apply this tactic. Consider the four goals: |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
697 |
*} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
698 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
699 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
700 |
lemma shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" |
104 | 701 |
apply(tactic {* SUBGOAL select_tac 4 *}) |
702 |
apply(tactic {* SUBGOAL select_tac 3 *}) |
|
703 |
apply(tactic {* SUBGOAL select_tac 2 *}) |
|
99 | 704 |
apply(tactic {* SUBGOAL select_tac 1 *}) |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
705 |
txt{* \begin{minipage}{\textwidth} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
706 |
@{subgoals [display]} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
707 |
\end{minipage} *} |
99 | 708 |
(*<*)oops(*>*) |
709 |
||
710 |
text {* |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
711 |
where in all but the last the tactic applied an introduction rule. |
109 | 712 |
Note that we applied the tactic to the goals in ``reverse'' order. |
713 |
This is a trick in order to be independent from the subgoals that are |
|
714 |
produced by the rule. If we had applied it in the other order |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
715 |
*} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
716 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
717 |
lemma shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
718 |
apply(tactic {* SUBGOAL select_tac 1 *}) |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
719 |
apply(tactic {* SUBGOAL select_tac 3 *}) |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
720 |
apply(tactic {* SUBGOAL select_tac 4 *}) |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
721 |
apply(tactic {* SUBGOAL select_tac 5 *}) |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
722 |
(*<*)oops(*>*) |
99 | 723 |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
724 |
text {* |
109 | 725 |
then we have to be careful to not apply the tactic to the two subgoals produced by |
726 |
the first goal. To do this can result in quite messy code. In contrast, |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
727 |
the ``reverse application'' is easy to implement. |
104 | 728 |
|
109 | 729 |
Of course, this example is contrived: there are much simpler methods available in |
730 |
Isabelle for implementing a proof procedure analysing a goal according to its topmost |
|
731 |
connective. These simpler methods use tactic combinators, which we will explain |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
732 |
in the next section. |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
733 |
*} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
734 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
735 |
section {* Tactic Combinators *} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
736 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
737 |
text {* |
109 | 738 |
The purpose of tactic combinators is to build compound tactics out of |
739 |
smaller tactics. In the previous section we already used @{ML THEN}, which |
|
740 |
just strings together two tactics in a sequence. For example: |
|
93 | 741 |
*} |
742 |
||
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
743 |
lemma shows "(Foo \<and> Bar) \<and> False" |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
744 |
apply(tactic {* rtac @{thm conjI} 1 THEN rtac @{thm conjI} 1 *}) |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
745 |
txt {* \begin{minipage}{\textwidth} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
746 |
@{subgoals [display]} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
747 |
\end{minipage} *} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
748 |
(*<*)oops(*>*) |
99 | 749 |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
750 |
text {* |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
751 |
If you want to avoid the hard-coded subgoal addressing, then you can use |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
752 |
the ``primed'' version of @{ML THEN}. For example: |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
753 |
*} |
93 | 754 |
|
99 | 755 |
lemma shows "(Foo \<and> Bar) \<and> False" |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
756 |
apply(tactic {* (rtac @{thm conjI} THEN' rtac @{thm conjI}) 1 *}) |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
757 |
txt {* \begin{minipage}{\textwidth} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
758 |
@{subgoals [display]} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
759 |
\end{minipage} *} |
93 | 760 |
(*<*)oops(*>*) |
761 |
||
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
762 |
text {* |
109 | 763 |
Here you only have to specify the subgoal of interest only once and |
764 |
it is consistently applied to the component tactics. |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
765 |
For most tactic combinators such a ``primed'' version exists and |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
766 |
in what follows we will usually prefer it over the ``unprimed'' one. |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
767 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
768 |
If there is a list of tactics that should all be tried out in |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
769 |
sequence, you can use the combinator @{ML EVERY'}. For example |
109 | 770 |
the function @{ML foo_tac'} from page~\pageref{tac:footacprime} can also |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
771 |
be written as: |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
772 |
*} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
773 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
774 |
ML{*val foo_tac'' = EVERY' [etac @{thm disjE}, rtac @{thm disjI2}, |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
775 |
atac, rtac @{thm disjI1}, atac]*} |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
776 |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
777 |
text {* |
109 | 778 |
There is even another way of implementing this tactic: in automatic proof |
779 |
procedures (in contrast to tactics that might be called by the user) there |
|
780 |
are often long lists of tactics that are applied to the first |
|
781 |
subgoal. Instead of writing the code above and then calling @{ML "foo_tac'' 1"}, |
|
782 |
you can also just write |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
783 |
*} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
784 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
785 |
ML{*val foo_tac1 = EVERY1 [etac @{thm disjE}, rtac @{thm disjI2}, |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
786 |
atac, rtac @{thm disjI1}, atac]*} |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
787 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
788 |
text {* |
118
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
789 |
and call @{ML foo_tac1}. |
109 | 790 |
|
791 |
With the combinators @{ML THEN'}, @{ML EVERY'} and @{ML EVERY1} it must be |
|
792 |
guaranteed that all component tactics successfully apply; otherwise the |
|
793 |
whole tactic will fail. If you rather want to try out a number of tactics, |
|
794 |
then you can use the combinator @{ML ORELSE'} for two tactics, and @{ML |
|
795 |
FIRST'} (or @{ML FIRST1}) for a list of tactics. For example, the tactic |
|
796 |
||
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
797 |
*} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
798 |
|
131 | 799 |
ML{*val orelse_xmp = rtac @{thm disjI1} ORELSE' rtac @{thm conjI}*} |
99 | 800 |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
801 |
text {* |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
802 |
will first try out whether rule @{text disjI} applies and after that |
109 | 803 |
@{text conjI}. To see this consider the proof |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
804 |
*} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
805 |
|
99 | 806 |
lemma shows "True \<and> False" and "Foo \<or> Bar" |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
807 |
apply(tactic {* orelse_xmp 2 *}) |
99 | 808 |
apply(tactic {* orelse_xmp 1 *}) |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
809 |
txt {* which results in the goal state |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
810 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
811 |
\begin{minipage}{\textwidth} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
812 |
@{subgoals [display]} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
813 |
\end{minipage} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
814 |
*} |
93 | 815 |
(*<*)oops(*>*) |
816 |
||
817 |
||
818 |
text {* |
|
109 | 819 |
Using @{ML FIRST'} we can simplify our @{ML select_tac} from Page~\pageref{tac:selecttac} |
820 |
as follows: |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
821 |
*} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
822 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
823 |
ML{*val select_tac' = FIRST' [rtac @{thm conjI}, rtac @{thm impI}, |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
824 |
rtac @{thm notI}, rtac @{thm allI}, K all_tac]*} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
825 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
826 |
text {* |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
827 |
Since we like to mimic the behaviour of @{ML select_tac} as closely as possible, |
109 | 828 |
we must include @{ML all_tac} at the end of the list, otherwise the tactic will |
118
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
829 |
fail if no rule applies (we also have to wrap @{ML all_tac} using the |
109 | 830 |
@{ML K}-combinator, because it does not take a subgoal number as argument). You can |
831 |
test the tactic on the same goals: |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
832 |
*} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
833 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
834 |
lemma shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
835 |
apply(tactic {* select_tac' 4 *}) |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
836 |
apply(tactic {* select_tac' 3 *}) |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
837 |
apply(tactic {* select_tac' 2 *}) |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
838 |
apply(tactic {* select_tac' 1 *}) |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
839 |
txt{* \begin{minipage}{\textwidth} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
840 |
@{subgoals [display]} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
841 |
\end{minipage} *} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
842 |
(*<*)oops(*>*) |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
843 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
844 |
text {* |
109 | 845 |
Since such repeated applications of a tactic to the reverse order of |
846 |
\emph{all} subgoals is quite common, there is the tactic combinator |
|
847 |
@{ML ALLGOALS} that simplifies this. Using this combinator you can simply |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
848 |
write: *} |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
849 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
850 |
lemma shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
851 |
apply(tactic {* ALLGOALS select_tac' *}) |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
852 |
txt{* \begin{minipage}{\textwidth} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
853 |
@{subgoals [display]} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
854 |
\end{minipage} *} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
855 |
(*<*)oops(*>*) |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
856 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
857 |
text {* |
109 | 858 |
Remember that we chose to implement @{ML select_tac'} so that it |
859 |
always succeeds. This can be potentially very confusing for the user, |
|
860 |
for example, in cases where the goal is the form |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
861 |
*} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
862 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
863 |
lemma shows "E \<Longrightarrow> F" |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
864 |
apply(tactic {* select_tac' 1 *}) |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
865 |
txt{* \begin{minipage}{\textwidth} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
866 |
@{subgoals [display]} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
867 |
\end{minipage} *} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
868 |
(*<*)oops(*>*) |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
869 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
870 |
text {* |
109 | 871 |
In this case no rule applies. The problem for the user is that there is little |
872 |
chance to see whether or not progress in the proof has been made. By convention |
|
873 |
therefore, tactics visible to the user should either change something or fail. |
|
874 |
||
875 |
To comply with this convention, we could simply delete the @{ML "K all_tac"} |
|
876 |
from the end of the theorem list. As a result @{ML select_tac'} would only |
|
877 |
succeed on goals where it can make progress. But for the sake of argument, |
|
878 |
let us suppose that this deletion is \emph{not} an option. In such cases, you can |
|
879 |
use the combinator @{ML CHANGED} to make sure the subgoal has been changed |
|
880 |
by the tactic. Because now |
|
881 |
||
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
882 |
*} |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
883 |
|
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
884 |
lemma shows "E \<Longrightarrow> F" |
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
885 |
apply(tactic {* CHANGED (select_tac' 1) *})(*<*)?(*>*) |
109 | 886 |
txt{* gives the error message: |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
887 |
\begin{isabelle} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
888 |
@{text "*** empty result sequence -- proof command failed"}\\ |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
889 |
@{text "*** At command \"apply\"."} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
890 |
\end{isabelle} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
891 |
*}(*<*)oops(*>*) |
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
892 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
893 |
|
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
894 |
text {* |
109 | 895 |
We can further extend @{ML select_tac'} so that it not just applies to the topmost |
896 |
connective, but also to the ones immediately ``underneath'', i.e.~analyse the goal |
|
897 |
completely. For this you can use the tactic combinator @{ML REPEAT}. As an example |
|
898 |
suppose the following tactic |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
899 |
*} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
900 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
901 |
ML{*val repeat_xmp = REPEAT (CHANGED (select_tac' 1)) *} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
902 |
|
109 | 903 |
text {* which applied to the proof *} |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
904 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
905 |
lemma shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)" |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
906 |
apply(tactic {* repeat_xmp *}) |
109 | 907 |
txt{* produces |
908 |
||
909 |
\begin{minipage}{\textwidth} |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
910 |
@{subgoals [display]} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
911 |
\end{minipage} *} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
912 |
(*<*)oops(*>*) |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
913 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
914 |
text {* |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
915 |
Here it is crucial that @{ML select_tac'} is prefixed with @{ML CHANGED}, |
109 | 916 |
because otherwise @{ML REPEAT} runs into an infinite loop (it applies the |
917 |
tactic as long as it succeeds). The function |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
918 |
@{ML REPEAT1} is similar, but runs the tactic at least once (failing if |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
919 |
this is not possible). |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
920 |
|
109 | 921 |
If you are after the ``primed'' version of @{ML repeat_xmp} then you |
922 |
need to implement it as |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
923 |
*} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
924 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
925 |
ML{*val repeat_xmp' = REPEAT o CHANGED o select_tac'*} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
926 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
927 |
text {* |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
928 |
since there are no ``primed'' versions of @{ML REPEAT} and @{ML CHANGED}. |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
929 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
930 |
If you look closely at the goal state above, the tactics @{ML repeat_xmp} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
931 |
and @{ML repeat_xmp'} are not yet quite what we are after: the problem is |
109 | 932 |
that goals 2 and 3 are not analysed. This is because the tactic |
933 |
is applied repeatedly only to the first subgoal. To analyse also all |
|
934 |
resulting subgoals, you can use the tactic combinator @{ML REPEAT_ALL_NEW}. |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
935 |
Suppose the tactic |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
936 |
*} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
937 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
938 |
ML{*val repeat_all_new_xmp = REPEAT_ALL_NEW (CHANGED o select_tac')*} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
939 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
940 |
text {* |
109 | 941 |
you see that the following goal |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
942 |
*} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
943 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
944 |
lemma shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)" |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
945 |
apply(tactic {* repeat_all_new_xmp 1 *}) |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
946 |
txt{* \begin{minipage}{\textwidth} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
947 |
@{subgoals [display]} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
948 |
\end{minipage} *} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
949 |
(*<*)oops(*>*) |
93 | 950 |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
951 |
text {* |
109 | 952 |
is completely analysed according to the theorems we chose to |
120
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
953 |
include in @{ML select_tac'}. |
109 | 954 |
|
955 |
Recall that tactics produce a lazy sequence of successor goal states. These |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
956 |
states can be explored using the command \isacommand{back}. For example |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
957 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
958 |
*} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
959 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
960 |
lemma "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
961 |
apply(tactic {* etac @{thm disjE} 1 *}) |
109 | 962 |
txt{* applies the rule to the first assumption yielding the goal state:\smallskip |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
963 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
964 |
\begin{minipage}{\textwidth} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
965 |
@{subgoals [display]} |
109 | 966 |
\end{minipage}\smallskip |
967 |
||
968 |
After typing |
|
969 |
*} |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
970 |
(*<*) |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
971 |
oops |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
972 |
lemma "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
973 |
apply(tactic {* etac @{thm disjE} 1 *}) |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
974 |
(*>*) |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
975 |
back |
109 | 976 |
txt{* the rule now applies to the second assumption.\smallskip |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
977 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
978 |
\begin{minipage}{\textwidth} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
979 |
@{subgoals [display]} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
980 |
\end{minipage} *} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
981 |
(*<*)oops(*>*) |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
982 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
983 |
text {* |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
984 |
Sometimes this leads to confusing behaviour of tactics and also has |
109 | 985 |
the potential to explode the search space for tactics. |
986 |
These problems can be avoided by prefixing the tactic with the tactic |
|
987 |
combinator @{ML DETERM}. |
|
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
988 |
*} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
989 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
990 |
lemma "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
991 |
apply(tactic {* DETERM (etac @{thm disjE} 1) *}) |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
992 |
txt {*\begin{minipage}{\textwidth} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
993 |
@{subgoals [display]} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
994 |
\end{minipage} *} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
995 |
(*<*)oops(*>*) |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
996 |
text {* |
118
5f003fdf2653
polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
114
diff
changeset
|
997 |
This combinator will prune the search space to just the first successful application. |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
998 |
Attempting to apply \isacommand{back} in this goal states gives the |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
999 |
error message: |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
1000 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
1001 |
\begin{isabelle} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
1002 |
@{text "*** back: no alternatives"}\\ |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
1003 |
@{text "*** At command \"back\"."} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
1004 |
\end{isabelle} |
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
1005 |
|
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
1006 |
\begin{readmore} |
109 | 1007 |
Most tactic combinators described in this section are defined in @{ML_file "Pure/tctical.ML"}. |
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents:
107
diff
changeset
|
1008 |
\end{readmore} |
107
258ce361ba1b
polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents:
105
diff
changeset
|
1009 |
|
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
1010 |
*} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
1011 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
1012 |
section {* Rewriting and Simplifier Tactics *} |
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
1013 |
|
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
1014 |
text {* |
93 | 1015 |
@{ML rewrite_goals_tac} |
120
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
1016 |
|
93 | 1017 |
@{ML ObjectLogic.full_atomize_tac} |
120
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
1018 |
|
93 | 1019 |
@{ML ObjectLogic.rulify_tac} |
120
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
1020 |
|
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1021 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1022 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1023 |
section {* Simprocs *} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1024 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1025 |
text {* |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1026 |
In Isabelle you can also implement custom simplification procedures, called |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1027 |
\emph{simprocs}. Simprocs can be triggered on a specified term-pattern and |
134 | 1028 |
rewrite a term according to a theorem. They are useful in cases where |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1029 |
a rewriting rule must be produced on ``demand'' or when rewriting by |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1030 |
simplification is too unpredictable and potentially loops. |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1031 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1032 |
To see how simprocs work, let us first write a simproc that just prints out |
132 | 1033 |
the pattern which triggers it and otherwise does nothing. For this |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1034 |
you can use the function: |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1035 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1036 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1037 |
ML %linenosgray{*fun fail_sp_aux simpset redex = |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1038 |
let |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1039 |
val ctxt = Simplifier.the_context simpset |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1040 |
val _ = warning ("The redex: " ^ (str_of_cterm ctxt redex)) |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1041 |
in |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1042 |
NONE |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1043 |
end*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1044 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1045 |
text {* |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1046 |
This function takes a simpset and a redex (a @{ML_type cterm}) as |
132 | 1047 |
arguments. In Lines 3 and~4, we first extract the context from the given |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1048 |
simpset and then print out a message containing the redex. The function |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1049 |
returns @{ML NONE} (standing for an optional @{ML_type thm}) since at the |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1050 |
moment we are \emph{not} interested in actually rewriting anything. We want |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1051 |
that the simproc is triggered by the pattern @{term "Suc n"}. This can be |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1052 |
done by adding the simproc to the current simproc as follows |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1053 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1054 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1055 |
simproc_setup fail_sp ("Suc n") = {* K fail_sp_aux *} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1056 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1057 |
text {* |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1058 |
where the second argument specifies the pattern and the right-hand side |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1059 |
contains the code of the simproc (we have to use @{ML K} since we ignoring |
131 | 1060 |
an argument about morphisms\footnote{FIXME: what does the morphism do?}). |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1061 |
After this, the simplifier is aware of the simproc and you can test whether |
131 | 1062 |
it fires on the lemma: |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1063 |
*} |
120
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
1064 |
|
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1065 |
lemma shows "Suc 0 = 1" |
131 | 1066 |
apply(simp) |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1067 |
(*<*)oops(*>*) |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1068 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1069 |
text {* |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1070 |
This will print out the message twice: once for the left-hand side and |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1071 |
once for the right-hand side. The reason is that during simplification the |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1072 |
simplifier will at some point reduce the term @{term "1::nat"} to @{term "Suc |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1073 |
0"}, and then the simproc ``fires'' also on that term. |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1074 |
|
131 | 1075 |
We can add or delete the simproc from the current simpset by the usual |
132 | 1076 |
\isacommand{declare}-statement. For example the simproc will be deleted |
1077 |
with the declaration |
|
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1078 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1079 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1080 |
declare [[simproc del: fail_sp]] |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1081 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1082 |
text {* |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1083 |
If you want to see what happens with just \emph{this} simproc, without any |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1084 |
interference from other rewrite rules, you can call @{text fail_sp} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1085 |
as follows: |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1086 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1087 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1088 |
lemma shows "Suc 0 = 1" |
132 | 1089 |
apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [@{simproc fail_sp}]) 1*}) |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1090 |
(*<*)oops(*>*) |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1091 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1092 |
text {* |
131 | 1093 |
Now the message shows up only once since the term @{term "1::nat"} is |
1094 |
left unchanged. |
|
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1095 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1096 |
Setting up a simproc using the command \isacommand{setup\_simproc} will |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1097 |
always add automatically the simproc to the current simpset. If you do not |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1098 |
want this, then you have to use a slightly different method for setting |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1099 |
up the simproc. First the function @{ML fail_sp_aux} needs to be modified |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1100 |
to |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1101 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1102 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1103 |
ML{*fun fail_sp_aux' simpset redex = |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1104 |
let |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1105 |
val ctxt = Simplifier.the_context simpset |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1106 |
val _ = warning ("The redex: " ^ (Syntax.string_of_term ctxt redex)) |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1107 |
in |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1108 |
NONE |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1109 |
end*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1110 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1111 |
text {* |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1112 |
Here the redex is given as a @{ML_type term}, instead of a @{ML_type cterm} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1113 |
(therefore we printing it out using the function @{ML string_of_term in Syntax}). |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1114 |
We can turn this function into a proper simproc using |
93 | 1115 |
*} |
1116 |
||
105
f49dc7e96235
added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents:
104
diff
changeset
|
1117 |
|
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1118 |
ML{*val fail_sp' = |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1119 |
let |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1120 |
val thy = @{theory} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1121 |
val pat = [@{term "Suc n"}] |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1122 |
in |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1123 |
Simplifier.simproc_i thy "fail_sp'" pat (K fail_sp_aux') |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1124 |
end*} |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1125 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1126 |
text {* |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1127 |
Here the pattern is given as @{ML_type term} (instead of @{ML_type cterm}). |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1128 |
The function also takes a list of patterns that can trigger the simproc. |
132 | 1129 |
Now the simproc is set up and can be explicitly added using |
1130 |
{ML addsimprocs} to a simpset whenerver |
|
1131 |
needed. |
|
1132 |
||
1133 |
Simprocs are applied from inside to outside and from left to right. You can |
|
1134 |
see this in the proof |
|
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1135 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1136 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1137 |
lemma shows "Suc (Suc 0) = (Suc 1)" |
131 | 1138 |
apply(tactic {* simp_tac (HOL_ss addsimprocs [fail_sp']) 1*}) |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1139 |
(*<*)oops(*>*) |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1140 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1141 |
text {* |
132 | 1142 |
The simproc @{ML fail_sp'} prints out the sequence |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1143 |
|
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1144 |
@{text [display] |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1145 |
"> Suc 0 |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1146 |
> Suc (Suc 0) |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1147 |
> Suc 1"} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1148 |
|
131 | 1149 |
To see how a simproc applies a theorem, let us implement a simproc that |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1150 |
rewrites terms according to the equation: |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1151 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1152 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1153 |
lemma plus_one: |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1154 |
shows "Suc n \<equiv> n + 1" by simp |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1155 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1156 |
text {* |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1157 |
Simprocs expect that the given equation is a meta-equation, however the |
131 | 1158 |
equation can contain preconditions (the simproc then will only fire if the |
132 | 1159 |
preconditions can be solved). To see that one has relatively precise control over |
131 | 1160 |
the rewriting with simprocs, let us further assume we want that the simproc |
1161 |
only rewrites terms ``greater'' than @{term "Suc 0"}. For this we can write |
|
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1162 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1163 |
|
131 | 1164 |
|
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1165 |
ML{*fun plus_one_sp_aux ss redex = |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1166 |
case redex of |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1167 |
@{term "Suc 0"} => NONE |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1168 |
| _ => SOME @{thm plus_one}*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1169 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1170 |
text {* |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1171 |
and set up the simproc as follows. |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1172 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1173 |
|
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1174 |
ML{*val plus_one_sp = |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1175 |
let |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1176 |
val thy = @{theory} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1177 |
val pat = [@{term "Suc n"}] |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1178 |
in |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1179 |
Simplifier.simproc_i thy "sproc +1" pat (K plus_one_sp_aux) |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1180 |
end*} |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1181 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1182 |
text {* |
132 | 1183 |
Now the simproc is set up so that it is triggered by terms |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1184 |
of the form @{term "Suc n"}, but inside the simproc we only produce |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1185 |
a theorem if the term is not @{term "Suc 0"}. The result you can see |
131 | 1186 |
in the following proof |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1187 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1188 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1189 |
lemma shows "P (Suc (Suc (Suc 0))) (Suc 0)" |
131 | 1190 |
apply(tactic {* simp_tac (HOL_ss addsimprocs [plus_one_sp]) 1*}) |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1191 |
txt{* |
131 | 1192 |
where the simproc produces the goal state |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1193 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1194 |
@{subgoals[display]} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1195 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1196 |
(*<*)oops(*>*) |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1197 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1198 |
text {* |
133
3e94ccc0f31e
polishing and start of the section about attributes
Christian Urban <urbanc@in.tum.de>
parents:
132
diff
changeset
|
1199 |
As usual with rewriting you have to worry about looping: you already have |
132 | 1200 |
a loop with @{ML plus_one_sp}, if you apply it with the default simpset (because |
1201 |
the default simpset contains a rule which just does the opposite of @{ML plus_one_sp}, |
|
1202 |
namely rewriting @{text [quotes] "+ 1"} to a successor). So you have to be careful |
|
1203 |
in choosing the right simpset to which you add a simproc. |
|
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1204 |
|
132 | 1205 |
Next let us implement a simproc that replaces terms of the form @{term "Suc n"} |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1206 |
with the number @{text n} increase by one. First we implement a function that |
132 | 1207 |
takes a term and produces the corresponding integer value. |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1208 |
|
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1209 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1210 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1211 |
ML{*fun dest_suc_trm ((Const (@{const_name "Suc"}, _)) $ t) = 1 + dest_suc_trm t |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1212 |
| dest_suc_trm t = snd (HOLogic.dest_number t)*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1213 |
|
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1214 |
text {* |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1215 |
It uses the library function @{ML dest_number in HOLogic} that transforms |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1216 |
(Isabelle) terms, like @{term "0::nat"}, @{term "1::nat"}, @{term "2::nat"} and so |
131 | 1217 |
on, into integer values. This function raises the exception @{ML TERM}, if |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1218 |
the term is not a number. The next function expects a pair consisting of a term |
131 | 1219 |
@{text t} (containing @{term Suc}s) and the corresponding integer value @{text n}. |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1220 |
*} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1221 |
|
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1222 |
ML %linenosgray{*fun get_thm ctxt (t, n) = |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1223 |
let |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1224 |
val num = HOLogic.mk_number @{typ "nat"} n |
132 | 1225 |
val goal = Logic.mk_equals (t, num) |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1226 |
in |
132 | 1227 |
Goal.prove ctxt [] [] goal (K (arith_tac ctxt 1)) |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1228 |
end*} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1229 |
|
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1230 |
text {* |
132 | 1231 |
From the integer value it generates the corresponding number term, called |
1232 |
@{text num} (Line 3), and then generates the meta-equation @{text "t \<equiv> num"} |
|
1233 |
(Line 4), which it proves by the arithmetic tactic in Line 6. |
|
1234 |
||
134 | 1235 |
For our purpose at the moment, proving the meta-equation using @{ML arith_tac} is |
132 | 1236 |
fine, but there is also an alternative employing the simplifier with a very |
1237 |
restricted simpset. For the kind of lemmas we want to prove, the simpset |
|
1238 |
@{text "num_ss"} in the code will suffice. |
|
1239 |
*} |
|
131 | 1240 |
|
132 | 1241 |
ML{*fun get_thm_alt ctxt (t, n) = |
1242 |
let |
|
1243 |
val num = HOLogic.mk_number @{typ "nat"} n |
|
1244 |
val goal = Logic.mk_equals (t, num) |
|
1245 |
val num_ss = HOL_ss addsimps [@{thm One_nat_def}, @{thm Let_def}] @ |
|
1246 |
@{thms nat_number} @ @{thms neg_simps} @ @{thms plus_nat.simps} |
|
1247 |
in |
|
1248 |
Goal.prove ctxt [] [] goal (K (simp_tac num_ss 1)) |
|
1249 |
end*} |
|
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1250 |
|
132 | 1251 |
text {* |
1252 |
The advantage of @{ML get_thm_alt} is that it leaves very little room for |
|
1253 |
something to go wrong; in contrast it is much more difficult to predict |
|
1254 |
what happens with @{ML arith_tac}, especially in more complicated |
|
1255 |
circumstances. The disatvantage of @{ML get_thm_alt} is to find a simpset |
|
1256 |
that is sufficiently powerful to solve every instance of the lemmas |
|
1257 |
we like to prove. This requires careful tuning, but is often necessary in |
|
1258 |
``production code''.\footnote{It would be of great help if there is another |
|
1259 |
way than tracing the simplifier to obtain the lemmas that are successfully |
|
1260 |
applied during simplification. Alas, there is none.} |
|
1261 |
||
1262 |
Anyway, either version can be used in the function that produces the actual |
|
1263 |
theorem for the simproc. |
|
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1264 |
*} |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1265 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1266 |
ML{*fun nat_number_sp_aux ss t = |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1267 |
let |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1268 |
val ctxt = Simplifier.the_context ss |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1269 |
in |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1270 |
SOME (get_thm ctxt (t, dest_suc_trm t)) |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1271 |
handle TERM _ => NONE |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1272 |
end*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1273 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1274 |
text {* |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1275 |
This function uses the fact that @{ML dest_suc_trm} might throw an exception |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1276 |
@{ML TERM}. In this case there is nothing that can be rewritten and therefore no |
131 | 1277 |
theorem is produced (i.e.~the function returns @{ML NONE}). To try out the simproc |
1278 |
on an example, you can set it up as follows: |
|
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1279 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1280 |
|
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1281 |
ML{*val nat_number_sp = |
132 | 1282 |
let |
1283 |
val thy = @{theory} |
|
1284 |
val pat = [@{term "Suc n"}] |
|
1285 |
in |
|
1286 |
Simplifier.simproc_i thy "nat_number" pat (K nat_number_sp_aux) |
|
1287 |
end*} |
|
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1288 |
|
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1289 |
text {* |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1290 |
Now in the lemma |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1291 |
*} |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1292 |
|
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1293 |
lemma "P (Suc (Suc 2)) (Suc 99) (0::nat) (Suc 4 + Suc 0) (Suc (0 + 0))" |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1294 |
apply(tactic {* simp_tac (HOL_ss addsimprocs [nat_number_sp]) 1*}) |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1295 |
txt {* |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1296 |
you obtain the more legible goal state |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1297 |
|
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1298 |
@{subgoals [display]} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1299 |
*} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1300 |
(*<*)oops(*>*) |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1301 |
|
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1302 |
text {* |
132 | 1303 |
where the simproc rewrites all @{term "Suc"}s except in the last argument. There it cannot |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1304 |
rewrite anything, because it does not know how to transform the term @{term "Suc (0 + 0)"} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1305 |
into a number. To solve this problem have a look at the next exercise. |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1306 |
|
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1307 |
\begin{exercise}\label{ex:addsimproc} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1308 |
Write a simproc that replaces terms of the form @{term "t\<^isub>1 + t\<^isub>2"} by their |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1309 |
result. You can assume the terms are ``proper'' numbers, that is of the form |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1310 |
@{term "0::nat"}, @{term "1::nat"}, @{term "2::nat"} and so on. |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1311 |
\end{exercise} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1312 |
|
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1313 |
(FIXME: We did not do anything with morphisms. Anything interesting |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1314 |
one can say about them?) |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
129
diff
changeset
|
1315 |
*} |
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1316 |
|
137 | 1317 |
section {* Conversions\label{sec:conversion} *} |
132 | 1318 |
|
135 | 1319 |
text {* |
139 | 1320 |
Conversions are meta-equalities depending on some input term. There type is |
1321 |
as follows: |
|
135 | 1322 |
*} |
1323 |
||
139 | 1324 |
ML {*type conv = Thm.cterm -> Thm.thm*} |
135 | 1325 |
|
1326 |
text {* |
|
139 | 1327 |
The simplest two conversions are @{ML "Conv.all_conv"} and @{ML "Conv.no_conv"}: |
135 | 1328 |
|
139 | 1329 |
@{ML_response_fake "Conv.all_conv @{cterm True}" "True \<equiv> True"} |
135 | 1330 |
|
139 | 1331 |
@{ML_response_fake "Conv.no_conv @{cterm True}" "*** Exception- CTERM (\"no conversion\", []) raised"} |
135 | 1332 |
|
139 | 1333 |
Further basic conversions are, for example, @{ML "Thm.beta_conversion"} and |
1334 |
@{ML "Conv.rewr_conv"}: |
|
135 | 1335 |
*} |
1336 |
||
139 | 1337 |
lemma true_conj1: "True \<and> P \<equiv> P" by simp |
135 | 1338 |
|
1339 |
text {* |
|
139 | 1340 |
@{ML_response_fake "Thm.beta_conversion true @{cterm \"(\<lambda>x. x \<or> False) True\"}" |
1341 |
"(\<lambda>x. x \<or> False) True \<equiv> True \<or> False"} |
|
1342 |
||
1343 |
@{ML_response_fake "Conv.rewr_conv @{thm true_conj1} @{cterm \"True \<or> False\"}" |
|
1344 |
"True \<and> False \<equiv> False"} |
|
1345 |
||
1346 |
Basic conversions can be combined with a number of conversionals, i.e. |
|
1347 |
conversion combinators: |
|
1348 |
||
1349 |
@{ML_response_fake "Conv.then_conv (Thm.beta_conversion true, Conv.rewr_conv @{thm true_conj1}) @{cterm \"(\<lambda>x. x \<and> False) True\"}" "(\<lambda>x. x \<and> False) True \<equiv> False"} |
|
1350 |
||
1351 |
@{ML_response_fake "Conv.else_conv (Conv.rewr_conv @{thm true_conj1}, Conv.all_conv) @{cterm \"P \<or> (True \<and> Q)\"}" "P \<or> (True \<and> Q) \<equiv> P \<or> (True \<and> Q)"} |
|
1352 |
||
1353 |
@{ML_response_fake "Conv.arg_conv (Conv.rewr_conv @{thm true_conj1}) @{cterm \"P \<or> (True \<and> Q)\"}" "P \<or> (True \<and> Q) \<equiv> P \<or> Q"} |
|
1354 |
||
1355 |
\begin{readmore} |
|
1356 |
See @{ML_file "Pure/conv.ML"} for more conversionals. Further basic conversions |
|
1357 |
can be found in, for example, @{ML_file "Pure/thm.ML"}, |
|
1358 |
@{ML_file "Pure/drule.ML"}, and @{ML_file "Pure/meta_simplifier.ML"}. |
|
1359 |
\end{readmore} |
|
1360 |
||
1361 |
Conversions are a thin layer on top of Isabelle's inference kernel, and may |
|
1362 |
be seen as a controllable, bare-bone version of Isabelle's simplifier. We |
|
1363 |
will demonstrate this feature in the following example. |
|
1364 |
||
1365 |
For the moment, let's assumes we want to simplify according to the rewrite rule |
|
1366 |
@{thm true_conj1}. As a conversion, this may look as follows: |
|
135 | 1367 |
*} |
1368 |
||
139 | 1369 |
ML {*fun all_true1_conv ctxt ct = |
1370 |
(case Thm.term_of ct of |
|
1371 |
@{term "op \<and>"} $ @{term True} $ _ => Conv.rewr_conv @{thm true_conj1} ct |
|
1372 |
| _ $ _ => Conv.combination_conv (all_true1_conv ctxt) (all_true1_conv ctxt) ct |
|
1373 |
| Abs _ => Conv.abs_conv (fn (_, cx) => all_true1_conv cx) ctxt ct |
|
1374 |
| _ => Conv.all_conv ct)*} |
|
1375 |
||
1376 |
text {* |
|
1377 |
Here is the new conversion in action: |
|
1378 |
||
1379 |
@{ML_response_fake "all_true1_conv @{context} @{cterm \"distinct [1, x] \<longrightarrow> True \<and> 1 \<noteq> x\"}" |
|
1380 |
"distinct [1, x] \<longrightarrow> True \<and> 1 \<noteq> x \<equiv> distinct [1, x] \<longrightarrow> 1 \<noteq> x"} |
|
1381 |
||
1382 |
Now, let's complicate the task: Rewrite according to the rule |
|
1383 |
@{thm true_conj1}, but only in the first arguments of @{term If}: |
|
135 | 1384 |
*} |
1385 |
||
139 | 1386 |
ML {*fun if_true1_conv ctxt ct = |
1387 |
(case Thm.term_of ct of |
|
1388 |
Const (@{const_name If}, _) $ _ => Conv.arg_conv (all_true1_conv ctxt) ct |
|
1389 |
| _ $ _ => Conv.combination_conv (if_true1_conv ctxt) (if_true1_conv ctxt) ct |
|
1390 |
| Abs _ => Conv.abs_conv (fn (_, cx) => if_true1_conv cx) ctxt ct |
|
1391 |
| _ => Conv.all_conv ct)*} |
|
135 | 1392 |
|
139 | 1393 |
text {* |
1394 |
Here is an application of this new conversion: |
|
1395 |
||
1396 |
@{ML_response_fake "if_true1_conv @{context} @{cterm \"if P (True \<and> 1 \<noteq> 2) then True \<and> True else True \<and> False\"}" "if P (True \<and> 1 \<noteq> 2) then True \<and> True else True \<and> False \<equiv> if P (1 \<noteq> 2) then True \<and> True else True \<and> False"} |
|
135 | 1397 |
*} |
1398 |
||
1399 |
text {* |
|
139 | 1400 |
Conversions can also be turned into a tactic with the @{ML CONVERSION} |
1401 |
tactical, and there are predefined conversionals to traverse a goal state |
|
1402 |
||
1403 |
@{term "\<And>x. P \<Longrightarrow> Q"}: |
|
1404 |
||
1405 |
The conversional |
|
1406 |
@{ML Conv.params_conv} applies a conversion to @{term "P \<Longrightarrow> Q"}, |
|
1407 |
@{ML Conv.prems_conv} applies a conversion to all premises in @{term "P \<Longrightarrow> Q"}, and |
|
1408 |
@{ML Conv.concl_conv} applies a conversion to the conclusion @{term Q} in @{term "P \<Longrightarrow> Q"}. |
|
1409 |
||
1410 |
Assume we want to apply @{ML all_true1_conv} only in the conclusion |
|
1411 |
of the goal, and @{ML if_true1_conv} should only be applied in the premises. |
|
1412 |
Here is a tactic doing exactly that: |
|
135 | 1413 |
*} |
1414 |
||
139 | 1415 |
ML {*local open Conv |
1416 |
in |
|
1417 |
val true1_tac = CSUBGOAL (fn (goal, i) => |
|
1418 |
let val ctxt = ProofContext.init (Thm.theory_of_cterm goal) |
|
1419 |
in |
|
1420 |
CONVERSION ( |
|
1421 |
Conv.params_conv ~1 (fn cx => |
|
1422 |
(Conv.prems_conv ~1 (if_true1_conv cx)) then_conv |
|
1423 |
(Conv.concl_conv ~1 (all_true1_conv cx))) ctxt) i |
|
1424 |
end) |
|
1425 |
end*} |
|
1426 |
||
1427 |
text {* FIXME: give example, show that the conditional rewriting works *} |
|
1428 |
||
135 | 1429 |
|
1430 |
text {* |
|
139 | 1431 |
Conversions are not restricted to work on certified terms only, they can also |
1432 |
be lifted to theorem transformers, i.e. functions mapping a theorem to a |
|
1433 |
theorem, by the help of @{ML Conv.fconv_rule}. As an example, consider the |
|
1434 |
conversion @{ML all_true1_conv} again: |
|
135 | 1435 |
|
139 | 1436 |
@{ML_response_fake "Conv.fconv_rule (all_true1_conv @{context}) @{lemma \"P \<or> (True \<and> \<not>P)\" by simp}" "P \<or> \<not>P"} |
135 | 1437 |
*} |
1438 |
||
1439 |
||
1440 |
||
1441 |
||
132 | 1442 |
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1443 |
section {* Structured Proofs *} |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1444 |
|
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1445 |
text {* TBD *} |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
128
diff
changeset
|
1446 |
|
95
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1447 |
lemma True |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1448 |
proof |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1449 |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1450 |
{ |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1451 |
fix A B C |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1452 |
assume r: "A & B \<Longrightarrow> C" |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1453 |
assume A B |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1454 |
then have "A & B" .. |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1455 |
then have C by (rule r) |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1456 |
} |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1457 |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1458 |
{ |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1459 |
fix A B C |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1460 |
assume r: "A & B \<Longrightarrow> C" |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1461 |
assume A B |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1462 |
note conjI [OF this] |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1463 |
note r [OF this] |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1464 |
} |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1465 |
oops |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1466 |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1467 |
ML {* fun prop ctxt s = |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1468 |
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
|
1469 |
|
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1470 |
ML {* |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1471 |
val ctxt0 = @{context}; |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1472 |
val ctxt = ctxt0; |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1473 |
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
|
1474 |
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
|
1475 |
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
|
1476 |
val this = [@{thm conjI} OF this]; |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1477 |
val this = r OF this; |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1478 |
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
|
1479 |
val this = Variable.export ctxt ctxt0 [this] |
7235374f34c8
added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents:
93
diff
changeset
|
1480 |
*} |
93 | 1481 |
|
1482 |
||
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
99
diff
changeset
|
1483 |
|
139 | 1484 |
end |