87 |
86 |
88 \index{tactic@ {\tt\slshape{}tactic}} |
87 \index{tactic@ {\tt\slshape{}tactic}} |
89 \index{raw\_tactic@ {\tt\slshape{}raw\_tactic}} |
88 \index{raw\_tactic@ {\tt\slshape{}raw\_tactic}} |
90 During the development of automatic proof procedures, you will often find it |
89 During the development of automatic proof procedures, you will often find it |
91 necessary to test a tactic on examples. This can be conveniently done with |
90 necessary to test a tactic on examples. This can be conveniently done with |
92 the command \isacommand{apply}@{text "(tactic \<verbopen> \<dots> \<verbclose>)"}. |
91 the command \isacommand{apply}\<open>(tactic \<verbopen> \<dots> \<verbclose>)\<close>. |
93 Consider the following sequence of tactics |
92 Consider the following sequence of tactics |
94 |
93 |
95 *} |
94 \<close> |
96 |
95 |
97 ML %grayML{*fun foo_tac ctxt = |
96 ML %grayML\<open>fun foo_tac ctxt = |
98 (eresolve_tac ctxt [@{thm disjE}] 1 |
97 (eresolve_tac ctxt [@{thm disjE}] 1 |
99 THEN resolve_tac ctxt [@{thm disjI2}] 1 |
98 THEN resolve_tac ctxt [@{thm disjI2}] 1 |
100 THEN assume_tac ctxt 1 |
99 THEN assume_tac ctxt 1 |
101 THEN resolve_tac ctxt [@{thm disjI1}] 1 |
100 THEN resolve_tac ctxt [@{thm disjI1}] 1 |
102 THEN assume_tac ctxt 1)*} |
101 THEN assume_tac ctxt 1)\<close> |
103 |
102 |
104 text {* and the Isabelle proof: *} |
103 text \<open>and the Isabelle proof:\<close> |
105 |
104 |
106 lemma |
105 lemma |
107 shows "P \<or> Q \<Longrightarrow> Q \<or> P" |
106 shows "P \<or> Q \<Longrightarrow> Q \<or> P" |
108 apply(tactic {* foo_tac @{context} *}) |
107 apply(tactic \<open>foo_tac @{context}\<close>) |
109 done |
108 done |
110 |
109 |
111 text {* |
110 text \<open> |
112 By using @{text "tactic \<verbopen> \<dots> \<verbclose>"} you can call from the |
111 By using \<open>tactic \<verbopen> \<dots> \<verbclose>\<close> you can call from the |
113 user-level of Isabelle the tactic @{ML foo_tac} or |
112 user-level of Isabelle the tactic @{ML foo_tac} or |
114 any other function that returns a tactic. There are some goal transformations |
113 any other function that returns a tactic. There are some goal transformations |
115 that are performed by @{text "tactic"}. They can be avoided by using |
114 that are performed by \<open>tactic\<close>. They can be avoided by using |
116 @{text "raw_tactic"} instead. |
115 \<open>raw_tactic\<close> instead. |
117 |
116 |
118 The tactic @{ML foo_tac} is just a sequence of simple tactics stringed |
117 The tactic @{ML foo_tac} is just a sequence of simple tactics stringed |
119 together by @{ML THEN}. As can be seen, each simple tactic in @{ML foo_tac} |
118 together by @{ML THEN}. As can be seen, each simple tactic in @{ML foo_tac} |
120 has a hard-coded number that stands for the subgoal analysed by the |
119 has a hard-coded number that stands for the subgoal analysed by the |
121 tactic (@{text "1"} stands for the first, or top-most, subgoal). This hard-coding |
120 tactic (\<open>1\<close> stands for the first, or top-most, subgoal). This hard-coding |
122 of goals is sometimes wanted, but usually it is not. To avoid the explicit numbering, |
121 of goals is sometimes wanted, but usually it is not. To avoid the explicit numbering, |
123 you can write |
122 you can write |
124 *} |
123 \<close> |
125 |
124 |
126 ML %grayML{*fun foo_tac' ctxt = |
125 ML %grayML\<open>fun foo_tac' ctxt = |
127 (eresolve_tac ctxt [@{thm disjE}] |
126 (eresolve_tac ctxt [@{thm disjE}] |
128 THEN' resolve_tac ctxt [@{thm disjI2}] |
127 THEN' resolve_tac ctxt [@{thm disjI2}] |
129 THEN' assume_tac ctxt |
128 THEN' assume_tac ctxt |
130 THEN' resolve_tac ctxt [@{thm disjI1}] |
129 THEN' resolve_tac ctxt [@{thm disjI1}] |
131 THEN' assume_tac ctxt)*}text_raw{*\label{tac:footacprime}*} |
130 THEN' assume_tac ctxt)\<close>text_raw\<open>\label{tac:footacprime}\<close> |
132 |
131 |
133 text {* |
132 text \<open> |
134 where @{ML_ind THEN' in Tactical} is used instead of @{ML THEN in |
133 where @{ML_ind THEN' in Tactical} is used instead of @{ML THEN in |
135 Tactical}. (For most combinators that combine tactics---@{ML THEN} is only |
134 Tactical}. (For most combinators that combine tactics---@{ML THEN} is only |
136 one such combinator---a ``primed'' version exists.) With @{ML foo_tac'} you |
135 one such combinator---a ``primed'' version exists.) With @{ML foo_tac'} you |
137 can give the number for the subgoal explicitly when the tactic is called. So |
136 can give the number for the subgoal explicitly when the tactic is called. So |
138 in the next proof you can first discharge the second subgoal, and |
137 in the next proof you can first discharge the second subgoal, and |
139 then the first. |
138 then the first. |
140 *} |
139 \<close> |
141 |
140 |
142 lemma |
141 lemma |
143 shows "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1" |
142 shows "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1" |
144 and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2" |
143 and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2" |
145 apply(tactic {* foo_tac' @{context} 2 *}) |
144 apply(tactic \<open>foo_tac' @{context} 2\<close>) |
146 apply(tactic {* foo_tac' @{context} 1 *}) |
145 apply(tactic \<open>foo_tac' @{context} 1\<close>) |
147 done |
146 done |
148 |
147 |
149 text {* |
148 text \<open> |
150 This kind of addressing is more difficult to achieve when the goal is |
149 This kind of addressing is more difficult to achieve when the goal is |
151 hard-coded inside the tactic. |
150 hard-coded inside the tactic. |
152 |
151 |
153 The tactics @{ML foo_tac} and @{ML foo_tac'} are specific |
152 The tactics @{ML foo_tac} and @{ML foo_tac'} are specific |
154 to goals of the form @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. If the goal is not of |
153 to goals of the form @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. If the goal is not of |
156 precise, tactics do not produce this error message; the message originates from the |
155 precise, tactics do not produce this error message; the message originates from the |
157 \isacommand{apply} wrapper that calls the tactic.} |
156 \isacommand{apply} wrapper that calls the tactic.} |
158 |
157 |
159 |
158 |
160 \begin{isabelle} |
159 \begin{isabelle} |
161 @{text "*** empty result sequence -- proof command failed"}\\ |
160 \<open>*** empty result sequence -- proof command failed\<close>\\ |
162 @{text "*** At command \"apply\"."} |
161 \<open>*** At command "apply".\<close> |
163 \end{isabelle} |
162 \end{isabelle} |
164 |
163 |
165 This means they failed. The reason for this error message is that tactics |
164 This means they failed. The reason for this error message is that tactics |
166 are functions mapping a goal state to a (lazy) sequence of successor |
165 are functions mapping a goal state to a (lazy) sequence of successor |
167 states. Hence the type of a tactic is: |
166 states. Hence the type of a tactic is: |
168 *} |
167 \<close> |
169 |
168 |
170 ML %grayML{*type tactic = thm -> thm Seq.seq*} |
169 ML %grayML\<open>type tactic = thm -> thm Seq.seq\<close> |
171 |
170 |
172 text {* |
171 text \<open> |
173 By convention, if a tactic fails, then it should return the empty sequence. |
172 By convention, if a tactic fails, then it should return the empty sequence. |
174 Therefore, your own tactics should not raise exceptions |
173 Therefore, your own tactics should not raise exceptions |
175 willy-nilly; only in very grave failure situations should a tactic raise the |
174 willy-nilly; only in very grave failure situations should a tactic raise the |
176 exception @{text THM}. |
175 exception \<open>THM\<close>. |
177 |
176 |
178 The simplest tactics are @{ML_ind no_tac in Tactical} and |
177 The simplest tactics are @{ML_ind no_tac in Tactical} and |
179 @{ML_ind all_tac in Tactical}. The first returns the empty sequence and |
178 @{ML_ind all_tac in Tactical}. The first returns the empty sequence and |
180 is defined as |
179 is defined as |
181 *} |
180 \<close> |
182 |
181 |
183 ML %grayML{*fun no_tac thm = Seq.empty*} |
182 ML %grayML\<open>fun no_tac thm = Seq.empty\<close> |
184 |
183 |
185 text {* |
184 text \<open> |
186 which means @{ML no_tac} always fails. The second returns the given theorem wrapped |
185 which means @{ML no_tac} always fails. The second returns the given theorem wrapped |
187 in a single member sequence; it is defined as |
186 in a single member sequence; it is defined as |
188 *} |
187 \<close> |
189 |
188 |
190 ML %grayML{*fun all_tac thm = Seq.single thm*} |
189 ML %grayML\<open>fun all_tac thm = Seq.single thm\<close> |
191 |
190 |
192 text {* |
191 text \<open> |
193 which means @{ML all_tac} always succeeds, but also does not make any progress |
192 which means @{ML all_tac} always succeeds, but also does not make any progress |
194 with the proof. |
193 with the proof. |
195 |
194 |
196 The lazy list of possible successor goal states shows through at the user-level |
195 The lazy list of possible successor goal states shows through at the user-level |
197 of Isabelle when using the command \isacommand{back}. For instance in the |
196 of Isabelle when using the command \isacommand{back}. For instance in the |
198 following proof there are two possibilities for how to apply |
197 following proof there are two possibilities for how to apply |
199 @{ML foo_tac'}: either using the first assumption or the second. |
198 @{ML foo_tac'}: either using the first assumption or the second. |
200 *} |
199 \<close> |
201 |
200 |
202 lemma |
201 lemma |
203 shows "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P" |
202 shows "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P" |
204 apply(tactic {* foo_tac' @{context} 1 *}) |
203 apply(tactic \<open>foo_tac' @{context} 1\<close>) |
205 back |
204 back |
206 done |
205 done |
207 |
206 |
208 |
207 |
209 text {* |
208 text \<open> |
210 By using \isacommand{back}, we construct the proof that uses the |
209 By using \isacommand{back}, we construct the proof that uses the |
211 second assumption. While in the proof above, it does not really matter which |
210 second assumption. While in the proof above, it does not really matter which |
212 assumption is used, in more interesting cases provability might depend |
211 assumption is used, in more interesting cases provability might depend |
213 on exploring different possibilities. |
212 on exploring different possibilities. |
214 |
213 |
223 one goal state to the next, are functions from theorems to theorem |
222 one goal state to the next, are functions from theorems to theorem |
224 (sequences). The surprise resolves by knowing that every |
223 (sequences). The surprise resolves by knowing that every |
225 goal state is indeed a theorem. To shed more light on this, |
224 goal state is indeed a theorem. To shed more light on this, |
226 let us modify the code of @{ML all_tac} to obtain the following |
225 let us modify the code of @{ML all_tac} to obtain the following |
227 tactic |
226 tactic |
228 *} |
227 \<close> |
229 |
228 |
230 ML %grayML{*fun my_print_tac ctxt thm = |
229 ML %grayML\<open>fun my_print_tac ctxt thm = |
231 let |
230 let |
232 val _ = tracing (Pretty.string_of (pretty_thm_no_vars ctxt thm)) |
231 val _ = tracing (Pretty.string_of (pretty_thm_no_vars ctxt thm)) |
233 in |
232 in |
234 Seq.single thm |
233 Seq.single thm |
235 end*} |
234 end\<close> |
236 |
235 |
237 text {* |
236 text \<open> |
238 which prints out the given theorem (using the string-function defined in |
237 which prints out the given theorem (using the string-function defined in |
239 Section~\ref{sec:printing}) and then behaves like @{ML all_tac}. With this |
238 Section~\ref{sec:printing}) and then behaves like @{ML all_tac}. With this |
240 tactic we are in the position to inspect every goal state in a proof. For |
239 tactic we are in the position to inspect every goal state in a proof. For |
241 this consider the proof in Figure~\ref{fig:goalstates}: as can be seen, |
240 this consider the proof in Figure~\ref{fig:goalstates}: as can be seen, |
242 internally every goal state is an implication of the form |
241 internally every goal state is an implication of the form |
243 |
242 |
244 @{text[display] "A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> #C"} |
243 @{text[display] "A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> #C"} |
245 |
244 |
246 where @{term C} is the goal to be proved and the @{term "A\<^sub>i"} are |
245 where @{term C} is the goal to be proved and the @{term "A\<^sub>i"} are |
247 the subgoals. So after setting up the lemma, the goal state is always of the |
246 the subgoals. So after setting up the lemma, the goal state is always of the |
248 form @{text "C \<Longrightarrow> #C"}; when the proof is finished we are left with @{text |
247 form \<open>C \<Longrightarrow> #C\<close>; when the proof is finished we are left with \<open>#C\<close>. Since the goal @{term C} can potentially be an implication, there is a |
249 "#C"}. Since the goal @{term C} can potentially be an implication, there is a |
|
250 ``protector'' wrapped around it (the wrapper is the outermost constant |
248 ``protector'' wrapped around it (the wrapper is the outermost constant |
251 @{text "Const (\"Pure.prop\", bool \<Rightarrow> bool)"}; in the figure we make it visible |
249 \<open>Const ("Pure.prop", bool \<Rightarrow> bool)\<close>; in the figure we make it visible |
252 as a @{text #}). This wrapper prevents that premises of @{text C} are |
250 as a \<open>#\<close>). This wrapper prevents that premises of \<open>C\<close> are |
253 misinterpreted as open subgoals. While tactics can operate on the subgoals |
251 misinterpreted as open subgoals. While tactics can operate on the subgoals |
254 (the @{text "A\<^sub>i"} above), they are expected to leave the conclusion |
252 (the \<open>A\<^sub>i\<close> above), they are expected to leave the conclusion |
255 @{term C} intact, with the exception of possibly instantiating schematic |
253 @{term C} intact, with the exception of possibly instantiating schematic |
256 variables. This instantiations of schematic variables can be observed |
254 variables. This instantiations of schematic variables can be observed |
257 on the user level. Have a look at the following definition and proof. |
255 on the user level. Have a look at the following definition and proof. |
258 *} |
256 \<close> |
259 |
257 |
260 text_raw {* |
258 text_raw \<open> |
261 \begin{figure}[p] |
259 \begin{figure}[p] |
262 \begin{boxedminipage}{\textwidth} |
260 \begin{boxedminipage}{\textwidth} |
263 \begin{isabelle} |
261 \begin{isabelle} |
264 *} |
262 \<close> |
265 notation (output) "Pure.prop" ("#_" [1000] 1000) |
263 notation (output) "Pure.prop" ("#_" [1000] 1000) |
266 |
264 |
267 lemma |
265 lemma |
268 shows "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B" |
266 shows "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B" |
269 apply(tactic {* my_print_tac @{context} *}) |
267 apply(tactic \<open>my_print_tac @{context}\<close>) |
270 |
268 |
271 txt{* \begin{minipage}{\textwidth} |
269 txt\<open>\begin{minipage}{\textwidth} |
272 @{subgoals [display]} |
270 @{subgoals [display]} |
273 \end{minipage}\medskip |
271 \end{minipage}\medskip |
274 |
272 |
275 \begin{minipage}{\textwidth} |
273 \begin{minipage}{\textwidth} |
276 \small\colorbox{gray!20}{ |
274 \small\colorbox{gray!20}{ |
277 \begin{tabular}{@ {}l@ {}} |
275 \begin{tabular}{@ {}l@ {}} |
278 internal goal state:\\ |
276 internal goal state:\\ |
279 @{raw_goal_state} |
277 @{raw_goal_state} |
280 \end{tabular}} |
278 \end{tabular}} |
281 \end{minipage}\medskip |
279 \end{minipage}\medskip |
282 *} |
280 \<close> |
283 |
281 |
284 apply(rule conjI) |
282 apply(rule conjI) |
285 apply(tactic {* my_print_tac @{context} *}) |
283 apply(tactic \<open>my_print_tac @{context}\<close>) |
286 |
284 |
287 txt{* \begin{minipage}{\textwidth} |
285 txt\<open>\begin{minipage}{\textwidth} |
288 @{subgoals [display]} |
286 @{subgoals [display]} |
289 \end{minipage}\medskip |
287 \end{minipage}\medskip |
290 |
288 |
291 \begin{minipage}{\textwidth} |
289 \begin{minipage}{\textwidth} |
292 \small\colorbox{gray!20}{ |
290 \small\colorbox{gray!20}{ |
293 \begin{tabular}{@ {}l@ {}} |
291 \begin{tabular}{@ {}l@ {}} |
294 internal goal state:\\ |
292 internal goal state:\\ |
295 @{raw_goal_state} |
293 @{raw_goal_state} |
296 \end{tabular}} |
294 \end{tabular}} |
297 \end{minipage}\medskip |
295 \end{minipage}\medskip |
298 *} |
296 \<close> |
299 |
297 |
300 apply(assumption) |
298 apply(assumption) |
301 apply(tactic {* my_print_tac @{context} *}) |
299 apply(tactic \<open>my_print_tac @{context}\<close>) |
302 |
300 |
303 txt{* \begin{minipage}{\textwidth} |
301 txt\<open>\begin{minipage}{\textwidth} |
304 @{subgoals [display]} |
302 @{subgoals [display]} |
305 \end{minipage}\medskip |
303 \end{minipage}\medskip |
306 |
304 |
307 \begin{minipage}{\textwidth} |
305 \begin{minipage}{\textwidth} |
308 \small\colorbox{gray!20}{ |
306 \small\colorbox{gray!20}{ |
309 \begin{tabular}{@ {}l@ {}} |
307 \begin{tabular}{@ {}l@ {}} |
310 internal goal state:\\ |
308 internal goal state:\\ |
311 @{raw_goal_state} |
309 @{raw_goal_state} |
312 \end{tabular}} |
310 \end{tabular}} |
313 \end{minipage}\medskip |
311 \end{minipage}\medskip |
314 *} |
312 \<close> |
315 |
313 |
316 apply(assumption) |
314 apply(assumption) |
317 apply(tactic {* my_print_tac @{context} *}) |
315 apply(tactic \<open>my_print_tac @{context}\<close>) |
318 |
316 |
319 txt{* \begin{minipage}{\textwidth} |
317 txt\<open>\begin{minipage}{\textwidth} |
320 @{subgoals [display]} |
318 @{subgoals [display]} |
321 \end{minipage}\medskip |
319 \end{minipage}\medskip |
322 |
320 |
323 \begin{minipage}{\textwidth} |
321 \begin{minipage}{\textwidth} |
324 \small\colorbox{gray!20}{ |
322 \small\colorbox{gray!20}{ |
325 \begin{tabular}{@ {}l@ {}} |
323 \begin{tabular}{@ {}l@ {}} |
326 internal goal state:\\ |
324 internal goal state:\\ |
327 @{raw_goal_state} |
325 @{raw_goal_state} |
328 \end{tabular}} |
326 \end{tabular}} |
329 \end{minipage}\medskip |
327 \end{minipage}\medskip |
330 *} |
328 \<close> |
331 (*<*)oops(*>*) |
329 (*<*)oops(*>*) |
332 text_raw {* |
330 text_raw \<open> |
333 \end{isabelle} |
331 \end{isabelle} |
334 \end{boxedminipage} |
332 \end{boxedminipage} |
335 \caption{The figure shows an Isabelle snippet of a proof where each |
333 \caption{The figure shows an Isabelle snippet of a proof where each |
336 intermediate goal state is printed by the Isabelle system and by @{ML |
334 intermediate goal state is printed by the Isabelle system and by @{ML |
337 my_print_tac}. The latter shows the goal state as represented internally |
335 my_print_tac}. The latter shows the goal state as represented internally |
338 (highlighted boxes). This tactic shows that every goal state in Isabelle is |
336 (highlighted boxes). This tactic shows that every goal state in Isabelle is |
339 represented by a theorem: when you start the proof of \mbox{@{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> |
337 represented by a theorem: when you start the proof of \mbox{\<open>\<lbrakk>A; B\<rbrakk> \<Longrightarrow> |
340 A \<and> B"}} the theorem is @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> #(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}; when |
338 A \<and> B\<close>} the theorem is \<open>(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> #(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)\<close>; when |
341 you finish the proof the theorem is @{text "#(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> |
339 you finish the proof the theorem is \<open>#(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> |
342 B)"}.\label{fig:goalstates}} |
340 B)\<close>.\label{fig:goalstates}} |
343 \end{figure} |
341 \end{figure} |
344 *} |
342 \<close> |
345 |
343 |
346 definition |
344 definition |
347 EQ_TRUE |
345 EQ_TRUE |
348 where |
346 where |
349 "EQ_TRUE X \<equiv> (X = True)" |
347 "EQ_TRUE X \<equiv> (X = True)" |
351 schematic_goal test: |
349 schematic_goal test: |
352 shows "EQ_TRUE ?X" |
350 shows "EQ_TRUE ?X" |
353 unfolding EQ_TRUE_def |
351 unfolding EQ_TRUE_def |
354 by (rule refl) |
352 by (rule refl) |
355 |
353 |
356 text {* |
354 text \<open> |
357 By using \isacommand{schematic\_lemma} it is possible to prove lemmas including |
355 By using \isacommand{schematic\_lemma} it is possible to prove lemmas including |
358 meta-variables on the user level. However, the proved theorem is not @{text "EQ_TRUE ?X"}, |
356 meta-variables on the user level. However, the proved theorem is not \<open>EQ_TRUE ?X\<close>, |
359 as one might expect, but @{thm test}. We can test this with: |
357 as one might expect, but @{thm test}. We can test this with: |
360 |
358 |
361 \begin{isabelle} |
359 \begin{isabelle} |
362 \isacommand{thm}~@{thm [source] test}\\ |
360 \isacommand{thm}~@{thm [source] test}\\ |
363 @{text ">"}~@{thm test} |
361 \<open>>\<close>~@{thm test} |
364 \end{isabelle} |
362 \end{isabelle} |
365 |
363 |
366 The reason for this result is that the schematic variable @{text "?X"} |
364 The reason for this result is that the schematic variable \<open>?X\<close> |
367 is instantiated inside the proof to be @{term True} and then the |
365 is instantiated inside the proof to be @{term True} and then the |
368 instantiation propagated to the ``outside''. |
366 instantiation propagated to the ``outside''. |
369 |
367 |
370 \begin{readmore} |
368 \begin{readmore} |
371 For more information about the internals of goals see \isccite{sec:tactical-goals}. |
369 For more information about the internals of goals see \isccite{sec:tactical-goals}. |
372 \end{readmore} |
370 \end{readmore} |
373 |
371 |
374 *} |
372 \<close> |
375 |
373 |
376 |
374 |
377 section {* Simple Tactics\label{sec:simpletacs} *} |
375 section \<open>Simple Tactics\label{sec:simpletacs}\<close> |
378 |
376 |
379 text {* |
377 text \<open> |
380 In this section we will introduce more of the simple tactics in Isabelle. The |
378 In this section we will introduce more of the simple tactics in Isabelle. The |
381 first one is @{ML_ind print_tac in Tactical}, which is quite useful |
379 first one is @{ML_ind print_tac in Tactical}, which is quite useful |
382 for low-level debugging of tactics. It just prints out a message and the current |
380 for low-level debugging of tactics. It just prints out a message and the current |
383 goal state. Unlike @{ML my_print_tac} shown earlier, it prints the goal state |
381 goal state. Unlike @{ML my_print_tac} shown earlier, it prints the goal state |
384 as the user would see it. For example, processing the proof |
382 as the user would see it. For example, processing the proof |
385 *} |
383 \<close> |
386 |
384 |
387 lemma |
385 lemma |
388 shows "False \<Longrightarrow> True" |
386 shows "False \<Longrightarrow> True" |
389 apply(tactic {* print_tac @{context} "foo message" *}) |
387 apply(tactic \<open>print_tac @{context} "foo message"\<close>) |
390 txt{*gives: |
388 txt\<open>gives: |
391 \begin{isabelle} |
389 \begin{isabelle} |
392 @{text "foo message"}\\[3mm] |
390 \<open>foo message\<close>\\[3mm] |
393 @{prop "False \<Longrightarrow> True"}\\ |
391 @{prop "False \<Longrightarrow> True"}\\ |
394 @{text " 1. False \<Longrightarrow> True"}\\ |
392 \<open> 1. False \<Longrightarrow> True\<close>\\ |
395 \end{isabelle} |
393 \end{isabelle} |
396 *} |
394 \<close> |
397 (*<*)oops(*>*) |
395 (*<*)oops(*>*) |
398 |
396 |
399 text {* |
397 text \<open> |
400 A simple tactic for easy discharge of any proof obligations, even difficult ones, is |
398 A simple tactic for easy discharge of any proof obligations, even difficult ones, is |
401 @{ML_ind cheat_tac in Skip_Proof} in the structure @{ML_struct Skip_Proof}. |
399 @{ML_ind cheat_tac in Skip_Proof} in the structure @{ML_struct Skip_Proof}. |
402 This tactic corresponds to the Isabelle command \isacommand{sorry} and is |
400 This tactic corresponds to the Isabelle command \isacommand{sorry} and is |
403 sometimes useful during the development of tactics. |
401 sometimes useful during the development of tactics. |
404 *} |
402 \<close> |
405 |
403 |
406 lemma |
404 lemma |
407 shows "False" and "Goldbach_conjecture" |
405 shows "False" and "Goldbach_conjecture" |
408 apply(tactic {* Skip_Proof.cheat_tac @{context} 1 *}) |
406 apply(tactic \<open>Skip_Proof.cheat_tac @{context} 1\<close>) |
409 txt{*\begin{minipage}{\textwidth} |
407 txt\<open>\begin{minipage}{\textwidth} |
410 @{subgoals [display]} |
408 @{subgoals [display]} |
411 \end{minipage}*} |
409 \end{minipage}\<close> |
412 (*<*)oops(*>*) |
410 (*<*)oops(*>*) |
413 |
411 |
414 text {* |
412 text \<open> |
415 Another simple tactic is the function @{ML_ind assume_tac in Tactic}, which, as shown |
413 Another simple tactic is the function @{ML_ind assume_tac in Tactic}, which, as shown |
416 earlier, corresponds to the assumption method. |
414 earlier, corresponds to the assumption method. |
417 *} |
415 \<close> |
418 |
416 |
419 lemma |
417 lemma |
420 shows "P \<Longrightarrow> P" |
418 shows "P \<Longrightarrow> P" |
421 apply(tactic {* assume_tac @{context} 1 *}) |
419 apply(tactic \<open>assume_tac @{context} 1\<close>) |
422 txt{*\begin{minipage}{\textwidth} |
420 txt\<open>\begin{minipage}{\textwidth} |
423 @{subgoals [display]} |
421 @{subgoals [display]} |
424 \end{minipage}*} |
422 \end{minipage}\<close> |
425 (*<*)oops(*>*) |
423 (*<*)oops(*>*) |
426 |
424 |
427 text {* |
425 text \<open> |
428 Similarly, @{ML_ind resolve_tac in Tactic}, @{ML_ind dresolve_tac in Tactic}, @{ML_ind eresolve_tac |
426 Similarly, @{ML_ind resolve_tac in Tactic}, @{ML_ind dresolve_tac in Tactic}, @{ML_ind eresolve_tac |
429 in Tactic} and @{ML_ind forward_tac in Tactic} correspond (roughly) to @{text |
427 in Tactic} and @{ML_ind forward_tac in Tactic} correspond (roughly) to \<open>rule\<close>, \<open>drule\<close>, \<open>erule\<close> and \<open>frule\<close>, respectively. Each of |
430 rule}, @{text drule}, @{text erule} and @{text frule}, respectively. Each of |
|
431 them takes a theorem as argument and attempts to apply it to a goal. Below |
428 them takes a theorem as argument and attempts to apply it to a goal. Below |
432 are three self-explanatory examples. |
429 are three self-explanatory examples. |
433 *} |
430 \<close> |
434 |
431 |
435 lemma |
432 lemma |
436 shows "P \<and> Q" |
433 shows "P \<and> Q" |
437 apply(tactic {* resolve_tac @{context} [@{thm conjI}] 1 *}) |
434 apply(tactic \<open>resolve_tac @{context} [@{thm conjI}] 1\<close>) |
438 txt{*\begin{minipage}{\textwidth} |
435 txt\<open>\begin{minipage}{\textwidth} |
439 @{subgoals [display]} |
436 @{subgoals [display]} |
440 \end{minipage}*} |
437 \end{minipage}\<close> |
441 (*<*)oops(*>*) |
438 (*<*)oops(*>*) |
442 |
439 |
443 lemma |
440 lemma |
444 shows "P \<and> Q \<Longrightarrow> False" |
441 shows "P \<and> Q \<Longrightarrow> False" |
445 apply(tactic {* eresolve_tac @{context} [@{thm conjE}] 1 *}) |
442 apply(tactic \<open>eresolve_tac @{context} [@{thm conjE}] 1\<close>) |
446 txt{*\begin{minipage}{\textwidth} |
443 txt\<open>\begin{minipage}{\textwidth} |
447 @{subgoals [display]} |
444 @{subgoals [display]} |
448 \end{minipage}*} |
445 \end{minipage}\<close> |
449 (*<*)oops(*>*) |
446 (*<*)oops(*>*) |
450 |
447 |
451 lemma |
448 lemma |
452 shows "False \<and> True \<Longrightarrow> False" |
449 shows "False \<and> True \<Longrightarrow> False" |
453 apply(tactic {* dresolve_tac @{context} [@{thm conjunct2}] 1 *}) |
450 apply(tactic \<open>dresolve_tac @{context} [@{thm conjunct2}] 1\<close>) |
454 txt{*\begin{minipage}{\textwidth} |
451 txt\<open>\begin{minipage}{\textwidth} |
455 @{subgoals [display]} |
452 @{subgoals [display]} |
456 \end{minipage}*} |
453 \end{minipage}\<close> |
457 (*<*)oops(*>*) |
454 (*<*)oops(*>*) |
458 |
455 |
459 text {* |
456 text \<open> |
460 The function @{ML_ind resolve_tac in Tactic} |
457 The function @{ML_ind resolve_tac in Tactic} |
461 expects a list of theorems as argument. From this list it will apply the |
458 expects a list of theorems as argument. From this list it will apply the |
462 first applicable theorem (later theorems that are also applicable can be |
459 first applicable theorem (later theorems that are also applicable can be |
463 explored via the lazy sequences mechanism). Given the code |
460 explored via the lazy sequences mechanism). Given the code |
464 *} |
461 \<close> |
465 |
462 |
466 ML %grayML{*fun resolve_xmp_tac ctxt = resolve_tac ctxt [@{thm impI}, @{thm conjI}]*} |
463 ML %grayML\<open>fun resolve_xmp_tac ctxt = resolve_tac ctxt [@{thm impI}, @{thm conjI}]\<close> |
467 |
464 |
468 text {* |
465 text \<open> |
469 an example for @{ML resolve_tac} is the following proof where first an outermost |
466 an example for @{ML resolve_tac} is the following proof where first an outermost |
470 implication is analysed and then an outermost conjunction. |
467 implication is analysed and then an outermost conjunction. |
471 *} |
468 \<close> |
472 |
469 |
473 lemma |
470 lemma |
474 shows "C \<longrightarrow> (A \<and> B)" |
471 shows "C \<longrightarrow> (A \<and> B)" |
475 and "(A \<longrightarrow> B) \<and> C" |
472 and "(A \<longrightarrow> B) \<and> C" |
476 apply(tactic {* resolve_xmp_tac @{context} 1 *}) |
473 apply(tactic \<open>resolve_xmp_tac @{context} 1\<close>) |
477 apply(tactic {* resolve_xmp_tac @{context} 2 *}) |
474 apply(tactic \<open>resolve_xmp_tac @{context} 2\<close>) |
478 txt{*\begin{minipage}{\textwidth} |
475 txt\<open>\begin{minipage}{\textwidth} |
479 @{subgoals [display]} |
476 @{subgoals [display]} |
480 \end{minipage}*} |
477 \end{minipage}\<close> |
481 (*<*)oops(*>*) |
478 (*<*)oops(*>*) |
482 |
479 |
483 text {* |
480 text \<open> |
484 Similar versions taking a list of theorems exist for the tactics |
481 Similar versions taking a list of theorems exist for the tactics |
485 @{ML_ind dresolve_tac in Tactic}, |
482 @{ML_ind dresolve_tac in Tactic}, |
486 @{ML_ind eresolve_tac in Tactic} and so on. |
483 @{ML_ind eresolve_tac in Tactic} and so on. |
487 |
484 |
488 Another simple tactic is @{ML_ind cut_facts_tac in Tactic}. It inserts a |
485 Another simple tactic is @{ML_ind cut_facts_tac in Tactic}. It inserts a |
489 list of theorems into the assumptions of the current goal state. Below we |
486 list of theorems into the assumptions of the current goal state. Below we |
490 will insert the definitions for the constants @{term True} and @{term |
487 will insert the definitions for the constants @{term True} and @{term |
491 False}. So |
488 False}. So |
492 *} |
489 \<close> |
493 |
490 |
494 lemma |
491 lemma |
495 shows "True \<noteq> False" |
492 shows "True \<noteq> False" |
496 apply(tactic {* cut_facts_tac [@{thm True_def}, @{thm False_def}] 1 *}) |
493 apply(tactic \<open>cut_facts_tac [@{thm True_def}, @{thm False_def}] 1\<close>) |
497 txt{*produces the goal state |
494 txt\<open>produces the goal state |
498 \begin{isabelle} |
495 \begin{isabelle} |
499 @{subgoals [display]} |
496 @{subgoals [display]} |
500 \end{isabelle}*} |
497 \end{isabelle}\<close> |
501 (*<*)oops(*>*) |
498 (*<*)oops(*>*) |
502 |
499 |
503 text {* |
500 text \<open> |
504 Often proofs on the ML-level involve elaborate operations on assumptions and |
501 Often proofs on the ML-level involve elaborate operations on assumptions and |
505 @{text "\<And>"}-quantified variables. To do such operations using the basic tactics |
502 \<open>\<And>\<close>-quantified variables. To do such operations using the basic tactics |
506 shown so far is very unwieldy and brittle. Some convenience and |
503 shown so far is very unwieldy and brittle. Some convenience and |
507 safety is provided by the functions @{ML_ind FOCUS in Subgoal} and |
504 safety is provided by the functions @{ML_ind FOCUS in Subgoal} and |
508 @{ML_ind SUBPROOF in Subgoal}. These tactics fix the parameters |
505 @{ML_ind SUBPROOF in Subgoal}. These tactics fix the parameters |
509 and bind the various components of a goal state to a record. |
506 and bind the various components of a goal state to a record. |
510 To see what happens, suppose the function defined in Figure~\ref{fig:sptac}, which |
507 To see what happens, suppose the function defined in Figure~\ref{fig:sptac}, which |
511 takes a record and just prints out the contents of this record. Then consider |
508 takes a record and just prints out the contents of this record. Then consider |
512 the proof: |
509 the proof: |
513 *} |
510 \<close> |
514 |
511 |
515 |
512 |
516 text_raw{* |
513 text_raw\<open> |
517 \begin{figure}[t] |
514 \begin{figure}[t] |
518 \begin{minipage}{\textwidth} |
515 \begin{minipage}{\textwidth} |
519 \begin{isabelle} |
516 \begin{isabelle} |
520 *} |
517 \<close> |
521 |
518 |
522 ML %grayML{*fun foc_tac {context, params, prems, asms, concl, schematics} = |
519 ML %grayML\<open>fun foc_tac {context, params, prems, asms, concl, schematics} = |
523 let |
520 let |
524 fun assgn1 f1 f2 xs = |
521 fun assgn1 f1 f2 xs = |
525 let |
522 let |
526 val out = map (fn (x, y) => Pretty.enum ":=" "" "" [f1 x, f2 y]) xs |
523 val out = map (fn (x, y) => Pretty.enum ":=" "" "" [f1 x, f2 y]) xs |
527 in |
524 in |
536 ("conclusion: ", pretty_cterm context concl), |
533 ("conclusion: ", pretty_cterm context concl), |
537 ("premises: ", pretty_thms_no_vars context prems), |
534 ("premises: ", pretty_thms_no_vars context prems), |
538 ("schematics: ", assgn2 (pretty_cterm context) (snd schematics))] |
535 ("schematics: ", assgn2 (pretty_cterm context) (snd schematics))] |
539 in |
536 in |
540 tracing (Pretty.string_of (Pretty.chunks pps)); all_tac |
537 tracing (Pretty.string_of (Pretty.chunks pps)); all_tac |
541 end*} |
538 end\<close> |
542 |
539 |
543 text_raw{* |
540 text_raw\<open> |
544 \end{isabelle} |
541 \end{isabelle} |
545 \end{minipage} |
542 \end{minipage} |
546 \caption{A function that prints out the various parameters provided by |
543 \caption{A function that prints out the various parameters provided by |
547 @{ML FOCUS in Subgoal} and @{ML SUBPROOF}. It uses the functions defined |
544 @{ML FOCUS in Subgoal} and @{ML SUBPROOF}. It uses the functions defined |
548 in Section~\ref{sec:printing} for extracting strings from @{ML_type cterm}s |
545 in Section~\ref{sec:printing} for extracting strings from @{ML_type cterm}s |
549 and @{ML_type thm}s.\label{fig:sptac}} |
546 and @{ML_type thm}s.\label{fig:sptac}} |
550 \end{figure} |
547 \end{figure} |
551 *} |
548 \<close> |
552 |
549 |
553 schematic_goal |
550 schematic_goal |
554 shows "\<And>x y. A x y \<Longrightarrow> B y x \<longrightarrow> C (?z y) x" |
551 shows "\<And>x y. A x y \<Longrightarrow> B y x \<longrightarrow> C (?z y) x" |
555 apply(tactic {* Subgoal.FOCUS foc_tac @{context} 1 *}) |
552 apply(tactic \<open>Subgoal.FOCUS foc_tac @{context} 1\<close>) |
556 |
553 |
557 txt {* |
554 txt \<open> |
558 The tactic produces the following printout: |
555 The tactic produces the following printout: |
559 |
556 |
560 \begin{quote}\small |
557 \begin{quote}\small |
561 \begin{tabular}{ll} |
558 \begin{tabular}{ll} |
562 params: & @{text "x:= x"}, @{text "y:= y"}\\ |
559 params: & \<open>x:= x\<close>, \<open>y:= y\<close>\\ |
563 assumptions: & @{term "A x y"}\\ |
560 assumptions: & @{term "A x y"}\\ |
564 conclusion: & @{term "B y x \<longrightarrow> C (z y) x"}\\ |
561 conclusion: & @{term "B y x \<longrightarrow> C (z y) x"}\\ |
565 premises: & @{term "A x y"}\\ |
562 premises: & @{term "A x y"}\\ |
566 schematics: & @{text "?z:=z"} |
563 schematics: & \<open>?z:=z\<close> |
567 \end{tabular} |
564 \end{tabular} |
568 \end{quote} |
565 \end{quote} |
569 |
566 |
570 The @{text params} and @{text schematics} stand for list of pairs where the |
567 The \<open>params\<close> and \<open>schematics\<close> stand for list of pairs where the |
571 left-hand side of @{text ":="} is replaced by the right-hand side inside the |
568 left-hand side of \<open>:=\<close> is replaced by the right-hand side inside the |
572 tactic. Notice that in the actual output the variables @{term x} and @{term |
569 tactic. Notice that in the actual output the variables @{term x} and @{term |
573 y} have a brown colour. Although they are parameters in the original goal, |
570 y} have a brown colour. Although they are parameters in the original goal, |
574 they are fixed inside the tactic. By convention these fixed variables are |
571 they are fixed inside the tactic. By convention these fixed variables are |
575 printed in brown colour. Similarly the schematic variable @{text ?z}. The |
572 printed in brown colour. Similarly the schematic variable \<open>?z\<close>. The |
576 assumption, or premise, @{prop "A x y"} is bound as @{ML_type cterm} to the |
573 assumption, or premise, @{prop "A x y"} is bound as @{ML_type cterm} to the |
577 record-variable @{text asms}, but also as @{ML_type thm} to @{text prems}. |
574 record-variable \<open>asms\<close>, but also as @{ML_type thm} to \<open>prems\<close>. |
578 |
575 |
579 If we continue the proof script by applying the @{text impI}-rule |
576 If we continue the proof script by applying the \<open>impI\<close>-rule |
580 *} |
577 \<close> |
581 |
578 |
582 apply(rule impI) |
579 apply(rule impI) |
583 apply(tactic {* Subgoal.FOCUS foc_tac @{context} 1 *}) |
580 apply(tactic \<open>Subgoal.FOCUS foc_tac @{context} 1\<close>) |
584 |
581 |
585 txt {* |
582 txt \<open> |
586 then the tactic prints out: |
583 then the tactic prints out: |
587 |
584 |
588 \begin{quote}\small |
585 \begin{quote}\small |
589 \begin{tabular}{ll} |
586 \begin{tabular}{ll} |
590 params: & @{text "x:= x"}, @{text "y:= y"}\\ |
587 params: & \<open>x:= x\<close>, \<open>y:= y\<close>\\ |
591 assumptions: & @{term "A x y"}, @{term "B y x"}\\ |
588 assumptions: & @{term "A x y"}, @{term "B y x"}\\ |
592 conclusion: & @{term "C (z y) x"}\\ |
589 conclusion: & @{term "C (z y) x"}\\ |
593 premises: & @{term "A x y"}, @{term "B y x"}\\ |
590 premises: & @{term "A x y"}, @{term "B y x"}\\ |
594 schematics: & @{text "?z:=z"} |
591 schematics: & \<open>?z:=z\<close> |
595 \end{tabular} |
592 \end{tabular} |
596 \end{quote} |
593 \end{quote} |
597 *} |
594 \<close> |
598 (*<*)oops(*>*) |
595 (*<*)oops(*>*) |
599 |
596 |
600 text {* |
597 text \<open> |
601 Now also @{term "B y x"} is an assumption bound to @{text asms} and @{text prems}. |
598 Now also @{term "B y x"} is an assumption bound to \<open>asms\<close> and \<open>prems\<close>. |
602 |
599 |
603 One difference between the tactics @{ML SUBPROOF} and @{ML FOCUS in Subgoal} |
600 One difference between the tactics @{ML SUBPROOF} and @{ML FOCUS in Subgoal} |
604 is that the former expects that the goal is solved completely, which the |
601 is that the former expects that the goal is solved completely, which the |
605 latter does not. Another is that @{ML SUBPROOF} cannot instantiate any schematic |
602 latter does not. Another is that @{ML SUBPROOF} cannot instantiate any schematic |
606 variables. |
603 variables. |
607 |
604 |
608 Observe that inside @{ML FOCUS in Subgoal} and @{ML SUBPROOF}, we are in a goal |
605 Observe that inside @{ML FOCUS in Subgoal} and @{ML SUBPROOF}, we are in a goal |
609 state where there is only a conclusion. This means the tactics @{ML dresolve_tac} and |
606 state where there is only a conclusion. This means the tactics @{ML dresolve_tac} and |
610 the like are of no use for manipulating the goal state. The assumptions inside |
607 the like are of no use for manipulating the goal state. The assumptions inside |
611 @{ML FOCUS in Subgoal} and @{ML SUBPROOF} are given as cterms and theorems in |
608 @{ML FOCUS in Subgoal} and @{ML SUBPROOF} are given as cterms and theorems in |
612 the arguments @{text "asms"} and @{text "prems"}, respectively. This |
609 the arguments \<open>asms\<close> and \<open>prems\<close>, respectively. This |
613 means we can apply them using the usual assumption tactics. With this you can |
610 means we can apply them using the usual assumption tactics. With this you can |
614 for example easily implement a tactic that behaves almost like @{ML assume_tac}: |
611 for example easily implement a tactic that behaves almost like @{ML assume_tac}: |
615 *} |
612 \<close> |
616 |
613 |
617 ML %grayML{*val atac' = Subgoal.FOCUS (fn {prems, context, ...} => resolve_tac context prems 1)*} |
614 ML %grayML\<open>val atac' = Subgoal.FOCUS (fn {prems, context, ...} => resolve_tac context prems 1)\<close> |
618 |
615 |
619 text {* |
616 text \<open> |
620 If you apply @{ML atac'} to the next lemma |
617 If you apply @{ML atac'} to the next lemma |
621 *} |
618 \<close> |
622 |
619 |
623 lemma |
620 lemma |
624 shows "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y" |
621 shows "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y" |
625 apply(tactic {* atac' @{context} 1 *}) |
622 apply(tactic \<open>atac' @{context} 1\<close>) |
626 txt{* it will produce |
623 txt\<open>it will produce |
627 @{subgoals [display]} *} |
624 @{subgoals [display]}\<close> |
628 (*<*)oops(*>*) |
625 (*<*)oops(*>*) |
629 |
626 |
630 text {* |
627 text \<open> |
631 Notice that @{ML atac'} inside @{ML FOCUS in Subgoal} calls @{ML |
628 Notice that @{ML atac'} inside @{ML FOCUS in Subgoal} calls @{ML |
632 resolve_tac} with the subgoal number @{text "1"} and also the outer call to |
629 resolve_tac} with the subgoal number \<open>1\<close> and also the outer call to |
633 @{ML FOCUS in Subgoal} in the \isacommand{apply}-step uses @{text "1"}. This |
630 @{ML FOCUS in Subgoal} in the \isacommand{apply}-step uses \<open>1\<close>. This |
634 is another advantage of @{ML FOCUS in Subgoal} and @{ML SUBPROOF}: the |
631 is another advantage of @{ML FOCUS in Subgoal} and @{ML SUBPROOF}: the |
635 addressing inside it is completely local to the tactic inside the |
632 addressing inside it is completely local to the tactic inside the |
636 subproof. It is therefore possible to also apply @{ML atac'} to the second |
633 subproof. It is therefore possible to also apply @{ML atac'} to the second |
637 goal by just writing: |
634 goal by just writing: |
638 |
635 |
639 *} |
636 \<close> |
640 |
637 |
641 lemma |
638 lemma |
642 shows "True" |
639 shows "True" |
643 and "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y" |
640 and "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y" |
644 apply(tactic {* atac' @{context} 2 *}) |
641 apply(tactic \<open>atac' @{context} 2\<close>) |
645 apply(rule TrueI) |
642 apply(rule TrueI) |
646 done |
643 done |
647 |
644 |
648 text {* |
645 text \<open> |
649 To sum up, both @{ML FOCUS in Subgoal} and @{ML SUBPROOF} are rather |
646 To sum up, both @{ML FOCUS in Subgoal} and @{ML SUBPROOF} are rather |
650 convenient, but can impose a considerable run-time penalty in automatic |
647 convenient, but can impose a considerable run-time penalty in automatic |
651 proofs. If speed is of the essence, then maybe the two lower level combinators |
648 proofs. If speed is of the essence, then maybe the two lower level combinators |
652 described next are more appropriate. |
649 described next are more appropriate. |
653 |
650 |
663 CSUBGOAL in Tactical}. They allow you to inspect a given subgoal (the former |
660 CSUBGOAL in Tactical}. They allow you to inspect a given subgoal (the former |
664 presents the subgoal as a @{ML_type term}, while the latter as a @{ML_type |
661 presents the subgoal as a @{ML_type term}, while the latter as a @{ML_type |
665 cterm}). With them you can implement a tactic that applies a rule according |
662 cterm}). With them you can implement a tactic that applies a rule according |
666 to the topmost logic connective in the subgoal (to illustrate this we only |
663 to the topmost logic connective in the subgoal (to illustrate this we only |
667 analyse a few connectives). The code of the tactic is as follows. |
664 analyse a few connectives). The code of the tactic is as follows. |
668 *} |
665 \<close> |
669 |
666 |
670 ML %linenosgray{*fun select_tac ctxt (t, i) = |
667 ML %linenosgray\<open>fun select_tac ctxt (t, i) = |
671 case t of |
668 case t of |
672 @{term "Trueprop"} $ t' => select_tac ctxt (t', i) |
669 @{term "Trueprop"} $ t' => select_tac ctxt (t', i) |
673 | @{term "(\<Longrightarrow>)"} $ _ $ t' => select_tac ctxt (t', i) |
670 | @{term "(\<Longrightarrow>)"} $ _ $ t' => select_tac ctxt (t', i) |
674 | @{term "(\<and>)"} $ _ $ _ => resolve_tac ctxt [@{thm conjI}] i |
671 | @{term "(\<and>)"} $ _ $ _ => resolve_tac ctxt [@{thm conjI}] i |
675 | @{term "(\<longrightarrow>)"} $ _ $ _ => resolve_tac ctxt [@{thm impI}] i |
672 | @{term "(\<longrightarrow>)"} $ _ $ _ => resolve_tac ctxt [@{thm impI}] i |
676 | @{term "Not"} $ _ => resolve_tac ctxt [@{thm notI}] i |
673 | @{term "Not"} $ _ => resolve_tac ctxt [@{thm notI}] i |
677 | Const (@{const_name "All"}, _) $ _ => resolve_tac ctxt [@{thm allI}] i |
674 | Const (@{const_name "All"}, _) $ _ => resolve_tac ctxt [@{thm allI}] i |
678 | _ => all_tac*}text_raw{*\label{tac:selecttac}*} |
675 | _ => all_tac\<close>text_raw\<open>\label{tac:selecttac}\<close> |
679 |
676 |
680 text {* |
677 text \<open> |
681 The input of the function is a term representing the subgoal and a number |
678 The input of the function is a term representing the subgoal and a number |
682 specifying the subgoal of interest. In Line 3 you need to descend under the |
679 specifying the subgoal of interest. In Line 3 you need to descend under the |
683 outermost @{term "Trueprop"} in order to get to the connective you like to |
680 outermost @{term "Trueprop"} in order to get to the connective you like to |
684 analyse. Otherwise goals like @{prop "A \<and> B"} are not properly |
681 analyse. Otherwise goals like @{prop "A \<and> B"} are not properly |
685 analysed. Similarly with meta-implications in the next line. While for the |
682 analysed. Similarly with meta-implications in the next line. While for the |
686 first five patterns we can use the @{text "@term"}-antiquotation to |
683 first five patterns we can use the \<open>@term\<close>-antiquotation to |
687 construct the patterns, the pattern in Line 8 cannot be constructed in this |
684 construct the patterns, the pattern in Line 8 cannot be constructed in this |
688 way. The reason is that an antiquotation would fix the type of the |
685 way. The reason is that an antiquotation would fix the type of the |
689 quantified variable. So you really have to construct this pattern using the |
686 quantified variable. So you really have to construct this pattern using the |
690 basic term-constructors. This is not necessary in the other cases, because their |
687 basic term-constructors. This is not necessary in the other cases, because their |
691 type is always fixed to function types involving only the type @{typ |
688 type is always fixed to function types involving only the type @{typ |
744 Since @{ML_ind resolve_tac in Tactic} and the like use higher-order unification, an |
741 Since @{ML_ind resolve_tac in Tactic} and the like use higher-order unification, an |
745 automatic proof procedure based on them might become too fragile, if it just |
742 automatic proof procedure based on them might become too fragile, if it just |
746 applies theorems as shown above. The reason is that a number of theorems |
743 applies theorems as shown above. The reason is that a number of theorems |
747 introduce schematic variables into the goal state. Consider for example the |
744 introduce schematic variables into the goal state. Consider for example the |
748 proof |
745 proof |
749 *} |
746 \<close> |
750 |
747 |
751 lemma |
748 lemma |
752 shows "\<forall>x \<in> A. P x \<Longrightarrow> Q x" |
749 shows "\<forall>x \<in> A. P x \<Longrightarrow> Q x" |
753 apply(tactic {* dresolve_tac @{context} [@{thm bspec}] 1 *}) |
750 apply(tactic \<open>dresolve_tac @{context} [@{thm bspec}] 1\<close>) |
754 txt{*\begin{minipage}{\textwidth} |
751 txt\<open>\begin{minipage}{\textwidth} |
755 @{subgoals [display]} |
752 @{subgoals [display]} |
756 \end{minipage}*} |
753 \end{minipage}\<close> |
757 (*<*)oops(*>*) |
754 (*<*)oops(*>*) |
758 |
755 |
759 text {* |
756 text \<open> |
760 where the application of theorem @{text bspec} generates two subgoals involving the |
757 where the application of theorem \<open>bspec\<close> generates two subgoals involving the |
761 new schematic variable @{text "?x"}. Now, if you are not careful, tactics |
758 new schematic variable \<open>?x\<close>. Now, if you are not careful, tactics |
762 applied to the first subgoal might instantiate this schematic variable in such a |
759 applied to the first subgoal might instantiate this schematic variable in such a |
763 way that the second subgoal becomes unprovable. If it is clear what the @{text "?x"} |
760 way that the second subgoal becomes unprovable. If it is clear what the \<open>?x\<close> |
764 should be, then this situation can be avoided by introducing a more |
761 should be, then this situation can be avoided by introducing a more |
765 constrained version of the @{text bspec}-theorem. One way to give such |
762 constrained version of the \<open>bspec\<close>-theorem. One way to give such |
766 constraints is by pre-instantiating theorems with other theorems. |
763 constraints is by pre-instantiating theorems with other theorems. |
767 The function @{ML_ind RS in Drule}, for example, does this. |
764 The function @{ML_ind RS in Drule}, for example, does this. |
768 |
765 |
769 @{ML_response_fake [display,gray] |
766 @{ML_response_fake [display,gray] |
770 "@{thm disjI1} RS @{thm conjI}" "\<lbrakk>?P1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<or> ?Q1) \<and> ?Q"} |
767 "@{thm disjI1} RS @{thm conjI}" "\<lbrakk>?P1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<or> ?Q1) \<and> ?Q"} |
771 |
768 |
772 In this example it instantiates the first premise of the @{text conjI}-theorem |
769 In this example it instantiates the first premise of the \<open>conjI\<close>-theorem |
773 with the theorem @{text disjI1}. If the instantiation is impossible, as in the |
770 with the theorem \<open>disjI1\<close>. If the instantiation is impossible, as in the |
774 case of |
771 case of |
775 |
772 |
776 @{ML_response_fake_both [display,gray] |
773 @{ML_response_fake_both [display,gray] |
777 "@{thm conjI} RS @{thm mp}" |
774 "@{thm conjI} RS @{thm mp}" |
778 "*** Exception- THM (\"RSN: no unifiers\", 1, |
775 "*** Exception- THM (\"RSN: no unifiers\", 1, |
815 congruence theorems. For example we would expect that the theorem |
812 congruence theorems. For example we would expect that the theorem |
816 @{thm [source] cong} |
813 @{thm [source] cong} |
817 |
814 |
818 \begin{isabelle} |
815 \begin{isabelle} |
819 \isacommand{thm}~@{thm [source] cong}\\ |
816 \isacommand{thm}~@{thm [source] cong}\\ |
820 @{text "> "}~@{thm cong} |
817 \<open>> \<close>~@{thm cong} |
821 \end{isabelle} |
818 \end{isabelle} |
822 |
819 |
823 is applicable in the following proof producing the subgoals |
820 is applicable in the following proof producing the subgoals |
824 @{term "t x = s u"} and @{term "y = w"}. |
821 @{term "t x = s u"} and @{term "y = w"}. |
825 *} |
822 \<close> |
826 |
823 |
827 lemma |
824 lemma |
828 fixes x y u w::"'a" |
825 fixes x y u w::"'a" |
829 shows "t x y = s u w" |
826 shows "t x y = s u w" |
830 apply(rule cong) |
827 apply(rule cong) |
831 txt{*\begin{minipage}{\textwidth} |
828 txt\<open>\begin{minipage}{\textwidth} |
832 @{subgoals [display]} |
829 @{subgoals [display]} |
833 \end{minipage}*} |
830 \end{minipage}\<close> |
834 (*<*)oops(*>*) |
831 (*<*)oops(*>*) |
835 |
832 |
836 text {* |
833 text \<open> |
837 As you can see this is unfortunately \emph{not} the case if we apply @{thm [source] |
834 As you can see this is unfortunately \emph{not} the case if we apply @{thm [source] |
838 cong} with the method @{text rule}. The problem is |
835 cong} with the method \<open>rule\<close>. The problem is |
839 that higher-order unification produces an instantiation that is not the |
836 that higher-order unification produces an instantiation that is not the |
840 intended one. While we can use \isacommand{back} to interactively find the |
837 intended one. While we can use \isacommand{back} to interactively find the |
841 intended instantiation, this is not an option for an automatic proof |
838 intended instantiation, this is not an option for an automatic proof |
842 procedure. Fortunately, the tactic @{ML_ind cong_tac in Cong_Tac} helps |
839 procedure. Fortunately, the tactic @{ML_ind cong_tac in Cong_Tac} helps |
843 with applying congruence theorems and finding the intended instantiation. |
840 with applying congruence theorems and finding the intended instantiation. |
844 For example |
841 For example |
845 *} |
842 \<close> |
846 |
843 |
847 lemma |
844 lemma |
848 fixes x y u w::"'a" |
845 fixes x y u w::"'a" |
849 shows "t x y = s u w" |
846 shows "t x y = s u w" |
850 apply(tactic {* Cong_Tac.cong_tac @{context} @{thm cong} 1 *}) |
847 apply(tactic \<open>Cong_Tac.cong_tac @{context} @{thm cong} 1\<close>) |
851 txt{*\begin{minipage}{\textwidth} |
848 txt\<open>\begin{minipage}{\textwidth} |
852 @{subgoals [display]} |
849 @{subgoals [display]} |
853 \end{minipage}*} |
850 \end{minipage}\<close> |
854 (*<*)oops(*>*) |
851 (*<*)oops(*>*) |
855 |
852 |
856 text {* |
853 text \<open> |
857 However, frequently it is necessary to explicitly match a theorem against a goal |
854 However, frequently it is necessary to explicitly match a theorem against a goal |
858 state and in doing so construct manually an appropriate instantiation. Imagine |
855 state and in doing so construct manually an appropriate instantiation. Imagine |
859 you have the theorem |
856 you have the theorem |
860 *} |
857 \<close> |
861 |
858 |
862 lemma rule: |
859 lemma rule: |
863 shows "\<lbrakk>f = g; x = y\<rbrakk> \<Longrightarrow> R (f x) (g y)" |
860 shows "\<lbrakk>f = g; x = y\<rbrakk> \<Longrightarrow> R (f x) (g y)" |
864 sorry |
861 sorry |
865 |
862 |
866 text {* |
863 text \<open> |
867 and you want to apply it to the goal @{term "(t\<^sub>1 t\<^sub>2) \<le> |
864 and you want to apply it to the goal @{term "(t\<^sub>1 t\<^sub>2) \<le> |
868 (s\<^sub>1 s\<^sub>2)"}. Since in the theorem all variables are |
865 (s\<^sub>1 s\<^sub>2)"}. Since in the theorem all variables are |
869 schematic, we have a nasty higher-order unification problem and @{text rtac} |
866 schematic, we have a nasty higher-order unification problem and \<open>rtac\<close> |
870 will not be able to apply this rule in the way we want. For the goal at hand |
867 will not be able to apply this rule in the way we want. For the goal at hand |
871 we want to use it so that @{term R} is instantiated to the constant @{text \<le>} and |
868 we want to use it so that @{term R} is instantiated to the constant \<open>\<le>\<close> and |
872 similarly ``obvious'' instantiations for the other variables. To achieve this |
869 similarly ``obvious'' instantiations for the other variables. To achieve this |
873 we need to match the conclusion of |
870 we need to match the conclusion of |
874 @{thm [source] rule} against the goal reading off an instantiation for |
871 @{thm [source] rule} against the goal reading off an instantiation for |
875 @{thm [source] rule}. For this the function @{ML_ind first_order_match in Thm} |
872 @{thm [source] rule}. For this the function @{ML_ind first_order_match in Thm} |
876 matches two @{ML_type cterm}s and produces, in the successful case, a matcher |
873 matches two @{ML_type cterm}s and produces, in the successful case, a matcher |
877 that can be used to instantiate the theorem. The instantiation |
874 that can be used to instantiate the theorem. The instantiation |
878 can be done with the function @{ML_ind instantiate_normalize in Drule}. To automate |
875 can be done with the function @{ML_ind instantiate_normalize in Drule}. To automate |
879 this we implement the following function. |
876 this we implement the following function. |
880 *} |
877 \<close> |
881 |
878 |
882 ML %linenosgray{*fun fo_rtac thm = Subgoal.FOCUS (fn {context, concl, ...} => |
879 ML %linenosgray\<open>fun fo_rtac thm = Subgoal.FOCUS (fn {context, concl, ...} => |
883 let |
880 let |
884 val concl_pat = Drule.strip_imp_concl (Thm.cprop_of thm) |
881 val concl_pat = Drule.strip_imp_concl (Thm.cprop_of thm) |
885 val insts = Thm.first_order_match (concl_pat, concl) |
882 val insts = Thm.first_order_match (concl_pat, concl) |
886 in |
883 in |
887 resolve_tac context [(Drule.instantiate_normalize insts thm)] 1 |
884 resolve_tac context [(Drule.instantiate_normalize insts thm)] 1 |
888 end)*} |
885 end)\<close> |
889 |
886 |
890 text {* |
887 text \<open> |
891 Note that we use @{ML FOCUS in Subgoal} because it gives us directly access |
888 Note that we use @{ML FOCUS in Subgoal} because it gives us directly access |
892 to the conclusion of the goal state, but also because this function |
889 to the conclusion of the goal state, but also because this function |
893 takes care of correctly handling parameters that might be present |
890 takes care of correctly handling parameters that might be present |
894 in the goal state. In Line 3 we use the function @{ML_ind strip_imp_concl in Drule} |
891 in the goal state. In Line 3 we use the function @{ML_ind strip_imp_concl in Drule} |
895 for calculating the conclusion of a theorem (produced as @{ML_type cterm}). |
892 for calculating the conclusion of a theorem (produced as @{ML_type cterm}). |
896 An example of @{ML fo_rtac} is as follows. |
893 An example of @{ML fo_rtac} is as follows. |
897 *} |
894 \<close> |
898 |
895 |
899 lemma |
896 lemma |
900 shows "\<And>t\<^sub>1 s\<^sub>1 (t\<^sub>2::'a) (s\<^sub>2::'a). (t\<^sub>1 t\<^sub>2) \<le> (s\<^sub>1 s\<^sub>2)" |
897 shows "\<And>t\<^sub>1 s\<^sub>1 (t\<^sub>2::'a) (s\<^sub>2::'a). (t\<^sub>1 t\<^sub>2) \<le> (s\<^sub>1 s\<^sub>2)" |
901 apply(tactic {* fo_rtac @{thm rule} @{context} 1 *}) |
898 apply(tactic \<open>fo_rtac @{thm rule} @{context} 1\<close>) |
902 txt{*\begin{minipage}{\textwidth} |
899 txt\<open>\begin{minipage}{\textwidth} |
903 @{subgoals [display]} |
900 @{subgoals [display]} |
904 \end{minipage}*} |
901 \end{minipage}\<close> |
905 (*<*)oops(*>*) |
902 (*<*)oops(*>*) |
906 |
903 |
907 text {* |
904 text \<open> |
908 We obtain the intended subgoals and also the parameters are correctly |
905 We obtain the intended subgoals and also the parameters are correctly |
909 introduced in both of them. Such manual instantiations are quite frequently |
906 introduced in both of them. Such manual instantiations are quite frequently |
910 necessary in order to appropriately constrain the application of theorems. |
907 necessary in order to appropriately constrain the application of theorems. |
911 Otherwise one can end up with ``wild'' higher-order unification problems |
908 Otherwise one can end up with ``wild'' higher-order unification problems |
912 that make automatic proofs fail. |
909 that make automatic proofs fail. |
914 \begin{readmore} |
911 \begin{readmore} |
915 Functions for matching @{ML_type cterm}s are defined in @{ML_file "Pure/thm.ML"}. |
912 Functions for matching @{ML_type cterm}s are defined in @{ML_file "Pure/thm.ML"}. |
916 Functions for instantiating schematic variables in theorems are defined |
913 Functions for instantiating schematic variables in theorems are defined |
917 in @{ML_file "Pure/drule.ML"}. |
914 in @{ML_file "Pure/drule.ML"}. |
918 \end{readmore} |
915 \end{readmore} |
919 *} |
916 \<close> |
920 |
917 |
921 section {* Tactic Combinators *} |
918 section \<open>Tactic Combinators\<close> |
922 |
919 |
923 text {* |
920 text \<open> |
924 The purpose of tactic combinators is to build compound tactics out of |
921 The purpose of tactic combinators is to build compound tactics out of |
925 smaller tactics. In the previous section we already used @{ML_ind THEN in Tactical}, |
922 smaller tactics. In the previous section we already used @{ML_ind THEN in Tactical}, |
926 which just strings together two tactics in a sequence. For example: |
923 which just strings together two tactics in a sequence. For example: |
927 *} |
924 \<close> |
928 |
925 |
929 lemma |
926 lemma |
930 shows "(Foo \<and> Bar) \<and> False" |
927 shows "(Foo \<and> Bar) \<and> False" |
931 apply(tactic {* resolve_tac @{context} [@{thm conjI}] 1 THEN resolve_tac @{context} [@{thm conjI}] 1 *}) |
928 apply(tactic \<open>resolve_tac @{context} [@{thm conjI}] 1 THEN resolve_tac @{context} [@{thm conjI}] 1\<close>) |
932 txt {* \begin{minipage}{\textwidth} |
929 txt \<open>\begin{minipage}{\textwidth} |
933 @{subgoals [display]} |
930 @{subgoals [display]} |
934 \end{minipage} *} |
931 \end{minipage}\<close> |
935 (*<*)oops(*>*) |
932 (*<*)oops(*>*) |
936 |
933 |
937 text {* |
934 text \<open> |
938 If you want to avoid the hard-coded subgoal addressing in each component, |
935 If you want to avoid the hard-coded subgoal addressing in each component, |
939 then, as seen earlier, you can use the ``primed'' version of @{ML THEN}. |
936 then, as seen earlier, you can use the ``primed'' version of @{ML THEN}. |
940 For example: |
937 For example: |
941 *} |
938 \<close> |
942 |
939 |
943 lemma |
940 lemma |
944 shows "(Foo \<and> Bar) \<and> False" |
941 shows "(Foo \<and> Bar) \<and> False" |
945 apply(tactic {* (resolve_tac @{context} [@{thm conjI}] THEN' resolve_tac @{context} [@{thm conjI}]) 1 *}) |
942 apply(tactic \<open>(resolve_tac @{context} [@{thm conjI}] THEN' resolve_tac @{context} [@{thm conjI}]) 1\<close>) |
946 txt {* \begin{minipage}{\textwidth} |
943 txt \<open>\begin{minipage}{\textwidth} |
947 @{subgoals [display]} |
944 @{subgoals [display]} |
948 \end{minipage} *} |
945 \end{minipage}\<close> |
949 (*<*)oops(*>*) |
946 (*<*)oops(*>*) |
950 |
947 |
951 text {* |
948 text \<open> |
952 Here you have to specify the subgoal of interest only once and it is |
949 Here you have to specify the subgoal of interest only once and it is |
953 consistently applied to the component. For most tactic combinators such a |
950 consistently applied to the component. For most tactic combinators such a |
954 ``primed'' version exists and in what follows we will usually prefer it over |
951 ``primed'' version exists and in what follows we will usually prefer it over |
955 the ``unprimed'' one. |
952 the ``unprimed'' one. |
956 |
953 |
957 The tactic combinator @{ML_ind RANGE in Tactical} takes a list of @{text n} |
954 The tactic combinator @{ML_ind RANGE in Tactical} takes a list of \<open>n\<close> |
958 tactics, say, and applies them to each of the first @{text n} subgoals. |
955 tactics, say, and applies them to each of the first \<open>n\<close> subgoals. |
959 For example below we first apply the introduction rule for conjunctions |
956 For example below we first apply the introduction rule for conjunctions |
960 and then apply a tactic to each of the two subgoals. |
957 and then apply a tactic to each of the two subgoals. |
961 *} |
958 \<close> |
962 |
959 |
963 lemma |
960 lemma |
964 shows "A \<Longrightarrow> True \<and> A" |
961 shows "A \<Longrightarrow> True \<and> A" |
965 apply(tactic {* (resolve_tac @{context} [@{thm conjI}] |
962 apply(tactic \<open>(resolve_tac @{context} [@{thm conjI}] |
966 THEN' RANGE [resolve_tac @{context} [@{thm TrueI}], assume_tac @{context}]) 1 *}) |
963 THEN' RANGE [resolve_tac @{context} [@{thm TrueI}], assume_tac @{context}]) 1\<close>) |
967 txt {* \begin{minipage}{\textwidth} |
964 txt \<open>\begin{minipage}{\textwidth} |
968 @{subgoals [display]} |
965 @{subgoals [display]} |
969 \end{minipage} *} |
966 \end{minipage}\<close> |
970 (*<*)oops(*>*) |
967 (*<*)oops(*>*) |
971 |
968 |
972 text {* |
969 text \<open> |
973 If there is a list of tactics that should all be tried out in sequence on |
970 If there is a list of tactics that should all be tried out in sequence on |
974 one specified subgoal, you can use the combinator @{ML_ind EVERY' in |
971 one specified subgoal, you can use the combinator @{ML_ind EVERY' in |
975 Tactical}. For example the function @{ML foo_tac'} from |
972 Tactical}. For example the function @{ML foo_tac'} from |
976 page~\pageref{tac:footacprime} can also be written as: |
973 page~\pageref{tac:footacprime} can also be written as: |
977 *} |
974 \<close> |
978 |
975 |
979 ML %grayML{*fun foo_tac'' ctxt = |
976 ML %grayML\<open>fun foo_tac'' ctxt = |
980 EVERY' [eresolve_tac ctxt [@{thm disjE}], resolve_tac ctxt [@{thm disjI2}], |
977 EVERY' [eresolve_tac ctxt [@{thm disjE}], resolve_tac ctxt [@{thm disjI2}], |
981 assume_tac ctxt, resolve_tac ctxt [@{thm disjI1}], assume_tac ctxt]*} |
978 assume_tac ctxt, resolve_tac ctxt [@{thm disjI1}], assume_tac ctxt]\<close> |
982 |
979 |
983 text {* |
980 text \<open> |
984 There is even another way of implementing this tactic: in automatic proof |
981 There is even another way of implementing this tactic: in automatic proof |
985 procedures (in contrast to tactics that might be called by the user) there |
982 procedures (in contrast to tactics that might be called by the user) there |
986 are often long lists of tactics that are applied to the first |
983 are often long lists of tactics that are applied to the first |
987 subgoal. Instead of writing the code above and then calling @{ML "foo_tac'' @{context} 1"}, |
984 subgoal. Instead of writing the code above and then calling @{ML "foo_tac'' @{context} 1"}, |
988 you can also just write |
985 you can also just write |
989 *} |
986 \<close> |
990 |
987 |
991 ML %grayML{*fun foo_tac1 ctxt = |
988 ML %grayML\<open>fun foo_tac1 ctxt = |
992 EVERY1 [eresolve_tac ctxt [@{thm disjE}], resolve_tac ctxt [@{thm disjI2}], |
989 EVERY1 [eresolve_tac ctxt [@{thm disjE}], resolve_tac ctxt [@{thm disjI2}], |
993 assume_tac ctxt, resolve_tac ctxt [@{thm disjI1}], assume_tac ctxt]*} |
990 assume_tac ctxt, resolve_tac ctxt [@{thm disjI1}], assume_tac ctxt]\<close> |
994 |
991 |
995 text {* |
992 text \<open> |
996 and call @{ML foo_tac1}. |
993 and call @{ML foo_tac1}. |
997 |
994 |
998 With the combinators @{ML THEN'}, @{ML EVERY'} and @{ML_ind EVERY1 in Tactical} it must be |
995 With the combinators @{ML THEN'}, @{ML EVERY'} and @{ML_ind EVERY1 in Tactical} it must be |
999 guaranteed that all component tactics successfully apply; otherwise the |
996 guaranteed that all component tactics successfully apply; otherwise the |
1000 whole tactic will fail. If you rather want to try out a number of tactics, |
997 whole tactic will fail. If you rather want to try out a number of tactics, |
1001 then you can use the combinator @{ML_ind ORELSE' in Tactical} for two tactics, and @{ML_ind |
998 then you can use the combinator @{ML_ind ORELSE' in Tactical} for two tactics, and @{ML_ind |
1002 FIRST' in Tactical} (or @{ML_ind FIRST1 in Tactical}) for a list of tactics. For example, the tactic |
999 FIRST' in Tactical} (or @{ML_ind FIRST1 in Tactical}) for a list of tactics. For example, the tactic |
1003 |
1000 |
1004 *} |
1001 \<close> |
1005 |
1002 |
1006 ML %grayML{*fun orelse_xmp_tac ctxt = |
1003 ML %grayML\<open>fun orelse_xmp_tac ctxt = |
1007 resolve_tac ctxt [@{thm disjI1}] ORELSE' resolve_tac ctxt [@{thm conjI}]*} |
1004 resolve_tac ctxt [@{thm disjI1}] ORELSE' resolve_tac ctxt [@{thm conjI}]\<close> |
1008 |
1005 |
1009 text {* |
1006 text \<open> |
1010 will first try out whether theorem @{text disjI} applies and in case of failure |
1007 will first try out whether theorem \<open>disjI\<close> applies and in case of failure |
1011 will try @{text conjI}. To see this consider the proof |
1008 will try \<open>conjI\<close>. To see this consider the proof |
1012 *} |
1009 \<close> |
1013 |
1010 |
1014 lemma |
1011 lemma |
1015 shows "True \<and> False" and "Foo \<or> Bar" |
1012 shows "True \<and> False" and "Foo \<or> Bar" |
1016 apply(tactic {* orelse_xmp_tac @{context} 2 *}) |
1013 apply(tactic \<open>orelse_xmp_tac @{context} 2\<close>) |
1017 apply(tactic {* orelse_xmp_tac @{context} 1 *}) |
1014 apply(tactic \<open>orelse_xmp_tac @{context} 1\<close>) |
1018 txt {* which results in the goal state |
1015 txt \<open>which results in the goal state |
1019 \begin{isabelle} |
1016 \begin{isabelle} |
1020 @{subgoals [display]} |
1017 @{subgoals [display]} |
1021 \end{isabelle} |
1018 \end{isabelle} |
1022 *} |
1019 \<close> |
1023 (*<*)oops(*>*) |
1020 (*<*)oops(*>*) |
1024 |
1021 |
1025 |
1022 |
1026 text {* |
1023 text \<open> |
1027 Using @{ML FIRST'} we can simplify our @{ML select_tac} from Page~\pageref{tac:selecttac} |
1024 Using @{ML FIRST'} we can simplify our @{ML select_tac} from Page~\pageref{tac:selecttac} |
1028 as follows: |
1025 as follows: |
1029 *} |
1026 \<close> |
1030 |
1027 |
1031 ML %grayML{*fun select_tac' ctxt = |
1028 ML %grayML\<open>fun select_tac' ctxt = |
1032 FIRST' [resolve_tac ctxt [@{thm conjI}], resolve_tac ctxt [@{thm impI}], |
1029 FIRST' [resolve_tac ctxt [@{thm conjI}], resolve_tac ctxt [@{thm impI}], |
1033 resolve_tac ctxt [@{thm notI}], resolve_tac ctxt [@{thm allI}], K all_tac]*}text_raw{*\label{tac:selectprime}*} |
1030 resolve_tac ctxt [@{thm notI}], resolve_tac ctxt [@{thm allI}], K all_tac]\<close>text_raw\<open>\label{tac:selectprime}\<close> |
1034 |
1031 |
1035 text {* |
1032 text \<open> |
1036 Since we like to mimic the behaviour of @{ML select_tac} as closely as possible, |
1033 Since we like to mimic the behaviour of @{ML select_tac} as closely as possible, |
1037 we must include @{ML all_tac} at the end of the list, otherwise the tactic will |
1034 we must include @{ML all_tac} at the end of the list, otherwise the tactic will |
1038 fail if no theorem applies (we also have to wrap @{ML all_tac} using the |
1035 fail if no theorem applies (we also have to wrap @{ML all_tac} using the |
1039 @{ML K}-combinator, because it does not take a subgoal number as argument). You can |
1036 @{ML K}-combinator, because it does not take a subgoal number as argument). You can |
1040 test the tactic on the same goals: |
1037 test the tactic on the same goals: |
1041 *} |
1038 \<close> |
1042 |
1039 |
1043 lemma |
1040 lemma |
1044 shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" |
1041 shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" |
1045 apply(tactic {* select_tac' @{context} 4 *}) |
1042 apply(tactic \<open>select_tac' @{context} 4\<close>) |
1046 apply(tactic {* select_tac' @{context} 3 *}) |
1043 apply(tactic \<open>select_tac' @{context} 3\<close>) |
1047 apply(tactic {* select_tac' @{context} 2 *}) |
1044 apply(tactic \<open>select_tac' @{context} 2\<close>) |
1048 apply(tactic {* select_tac' @{context} 1 *}) |
1045 apply(tactic \<open>select_tac' @{context} 1\<close>) |
1049 txt{* \begin{minipage}{\textwidth} |
1046 txt\<open>\begin{minipage}{\textwidth} |
1050 @{subgoals [display]} |
1047 @{subgoals [display]} |
1051 \end{minipage} *} |
1048 \end{minipage}\<close> |
1052 (*<*)oops(*>*) |
1049 (*<*)oops(*>*) |
1053 |
1050 |
1054 text {* |
1051 text \<open> |
1055 Since such repeated applications of a tactic to the reverse order of |
1052 Since such repeated applications of a tactic to the reverse order of |
1056 \emph{all} subgoals is quite common, there is the tactic combinator |
1053 \emph{all} subgoals is quite common, there is the tactic combinator |
1057 @{ML_ind ALLGOALS in Tactical} that simplifies this. Using this combinator you can simply |
1054 @{ML_ind ALLGOALS in Tactical} that simplifies this. Using this combinator you can simply |
1058 write: *} |
1055 write:\<close> |
1059 |
1056 |
1060 lemma |
1057 lemma |
1061 shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" |
1058 shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F" |
1062 apply(tactic {* ALLGOALS (select_tac' @{context}) *}) |
1059 apply(tactic \<open>ALLGOALS (select_tac' @{context})\<close>) |
1063 txt{* \begin{minipage}{\textwidth} |
1060 txt\<open>\begin{minipage}{\textwidth} |
1064 @{subgoals [display]} |
1061 @{subgoals [display]} |
1065 \end{minipage} *} |
1062 \end{minipage}\<close> |
1066 (*<*)oops(*>*) |
1063 (*<*)oops(*>*) |
1067 |
1064 |
1068 text {* |
1065 text \<open> |
1069 Remember that we chose to implement @{ML select_tac'} so that it |
1066 Remember that we chose to implement @{ML select_tac'} so that it |
1070 always succeeds by fact of having @{ML all_tac} at the end of the tactic |
1067 always succeeds by fact of having @{ML all_tac} at the end of the tactic |
1071 list. The same can be achieved with the tactic combinator @{ML_ind TRY in Tactical}. |
1068 list. The same can be achieved with the tactic combinator @{ML_ind TRY in Tactical}. |
1072 For example: |
1069 For example: |
1073 *} |
1070 \<close> |
1074 |
1071 |
1075 ML %grayML{*fun select_tac'' ctxt = |
1072 ML %grayML\<open>fun select_tac'' ctxt = |
1076 TRY o FIRST' [resolve_tac ctxt [@{thm conjI}], resolve_tac ctxt [@{thm impI}], |
1073 TRY o FIRST' [resolve_tac ctxt [@{thm conjI}], resolve_tac ctxt [@{thm impI}], |
1077 resolve_tac ctxt [@{thm notI}], resolve_tac ctxt [@{thm allI}]]*} |
1074 resolve_tac ctxt [@{thm notI}], resolve_tac ctxt [@{thm allI}]]\<close> |
1078 text_raw{*\label{tac:selectprimeprime}*} |
1075 text_raw\<open>\label{tac:selectprimeprime}\<close> |
1079 |
1076 |
1080 text {* |
1077 text \<open> |
1081 This tactic behaves in the same way as @{ML select_tac'}: it tries out |
1078 This tactic behaves in the same way as @{ML select_tac'}: it tries out |
1082 one of the given tactics and if none applies leaves the goal state |
1079 one of the given tactics and if none applies leaves the goal state |
1083 unchanged. This, however, can be potentially very confusing when visible to |
1080 unchanged. This, however, can be potentially very confusing when visible to |
1084 the user, for example, in cases where the goal is the form |
1081 the user, for example, in cases where the goal is the form |
1085 |
1082 |
1086 *} |
1083 \<close> |
1087 |
1084 |
1088 lemma |
1085 lemma |
1089 shows "E \<Longrightarrow> F" |
1086 shows "E \<Longrightarrow> F" |
1090 apply(tactic {* select_tac' @{context} 1 *}) |
1087 apply(tactic \<open>select_tac' @{context} 1\<close>) |
1091 txt{* \begin{minipage}{\textwidth} |
1088 txt\<open>\begin{minipage}{\textwidth} |
1092 @{subgoals [display]} |
1089 @{subgoals [display]} |
1093 \end{minipage} *} |
1090 \end{minipage}\<close> |
1094 (*<*)oops(*>*) |
1091 (*<*)oops(*>*) |
1095 |
1092 |
1096 text {* |
1093 text \<open> |
1097 In this case no theorem applies. But because we wrapped the tactic in a @{ML |
1094 In this case no theorem applies. But because we wrapped the tactic in a @{ML |
1098 TRY}, it does not fail anymore. The problem is that for the user there is |
1095 TRY}, it does not fail anymore. The problem is that for the user there is |
1099 little chance to see whether progress in the proof has been made, or not. By |
1096 little chance to see whether progress in the proof has been made, or not. By |
1100 convention therefore, tactics visible to the user should either change |
1097 convention therefore, tactics visible to the user should either change |
1101 something or fail. |
1098 something or fail. |
1104 in @{ML select_tac'} or delete @{ML TRY} from @{ML select_tac''}. But for |
1101 in @{ML select_tac'} or delete @{ML TRY} from @{ML select_tac''}. But for |
1105 the sake of argument, let us suppose that this deletion is \emph{not} an |
1102 the sake of argument, let us suppose that this deletion is \emph{not} an |
1106 option. In such cases, you can use the combinator @{ML_ind CHANGED in |
1103 option. In such cases, you can use the combinator @{ML_ind CHANGED in |
1107 Tactical} to make sure the subgoal has been changed by the tactic. Because |
1104 Tactical} to make sure the subgoal has been changed by the tactic. Because |
1108 now |
1105 now |
1109 *} |
1106 \<close> |
1110 |
1107 |
1111 lemma |
1108 lemma |
1112 shows "E \<Longrightarrow> F" |
1109 shows "E \<Longrightarrow> F" |
1113 apply(tactic {* CHANGED (select_tac' @{context} 1) *})(*<*)?(*>*) |
1110 apply(tactic \<open>CHANGED (select_tac' @{context} 1)\<close>)(*<*)?(*>*) |
1114 txt{* gives the error message: |
1111 txt\<open>gives the error message: |
1115 \begin{isabelle} |
1112 \begin{isabelle} |
1116 @{text "*** empty result sequence -- proof command failed"}\\ |
1113 \<open>*** empty result sequence -- proof command failed\<close>\\ |
1117 @{text "*** At command \"apply\"."} |
1114 \<open>*** At command "apply".\<close> |
1118 \end{isabelle} |
1115 \end{isabelle} |
1119 *}(*<*)oops(*>*) |
1116 \<close>(*<*)oops(*>*) |
1120 |
1117 |
1121 |
1118 |
1122 text {* |
1119 text \<open> |
1123 We can further extend the @{ML select_tac}s so that they not just apply to the topmost |
1120 We can further extend the @{ML select_tac}s so that they not just apply to the topmost |
1124 connective, but also to the ones immediately ``underneath'', i.e.~analyse the goal |
1121 connective, but also to the ones immediately ``underneath'', i.e.~analyse the goal |
1125 completely. For this you can use the tactic combinator @{ML_ind REPEAT in Tactical}. As an example |
1122 completely. For this you can use the tactic combinator @{ML_ind REPEAT in Tactical}. As an example |
1126 suppose the following tactic |
1123 suppose the following tactic |
1127 *} |
1124 \<close> |
1128 |
1125 |
1129 ML %grayML{*fun repeat_xmp_tac ctxt = REPEAT (CHANGED (select_tac' ctxt 1)) *} |
1126 ML %grayML\<open>fun repeat_xmp_tac ctxt = REPEAT (CHANGED (select_tac' ctxt 1))\<close> |
1130 |
1127 |
1131 text {* which applied to the proof *} |
1128 text \<open>which applied to the proof\<close> |
1132 |
1129 |
1133 lemma |
1130 lemma |
1134 shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)" |
1131 shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)" |
1135 apply(tactic {* repeat_xmp_tac @{context} *}) |
1132 apply(tactic \<open>repeat_xmp_tac @{context}\<close>) |
1136 txt{* produces |
1133 txt\<open>produces |
1137 \begin{isabelle} |
1134 \begin{isabelle} |
1138 @{subgoals [display]} |
1135 @{subgoals [display]} |
1139 \end{isabelle} *} |
1136 \end{isabelle}\<close> |
1140 (*<*)oops(*>*) |
1137 (*<*)oops(*>*) |
1141 |
1138 |
1142 text {* |
1139 text \<open> |
1143 Here it is crucial that @{ML select_tac'} is prefixed with @{ML CHANGED}, |
1140 Here it is crucial that @{ML select_tac'} is prefixed with @{ML CHANGED}, |
1144 because otherwise @{ML REPEAT} runs into an infinite loop (it applies the |
1141 because otherwise @{ML REPEAT} runs into an infinite loop (it applies the |
1145 tactic as long as it succeeds). The function @{ML_ind REPEAT1 in Tactical} |
1142 tactic as long as it succeeds). The function @{ML_ind REPEAT1 in Tactical} |
1146 is similar, but runs the tactic at least once (failing if this is not |
1143 is similar, but runs the tactic at least once (failing if this is not |
1147 possible). |
1144 possible). |
1148 |
1145 |
1149 If you are after the ``primed'' version of @{ML repeat_xmp_tac}, then you |
1146 If you are after the ``primed'' version of @{ML repeat_xmp_tac}, then you |
1150 can implement it as |
1147 can implement it as |
1151 *} |
1148 \<close> |
1152 |
1149 |
1153 ML %grayML{*fun repeat_xmp_tac' ctxt = REPEAT o CHANGED o select_tac' ctxt*} |
1150 ML %grayML\<open>fun repeat_xmp_tac' ctxt = REPEAT o CHANGED o select_tac' ctxt\<close> |
1154 |
1151 |
1155 text {* |
1152 text \<open> |
1156 since there are no ``primed'' versions of @{ML REPEAT} and @{ML CHANGED}. |
1153 since there are no ``primed'' versions of @{ML REPEAT} and @{ML CHANGED}. |
1157 |
1154 |
1158 If you look closely at the goal state above, then you see the tactics @{ML repeat_xmp_tac} |
1155 If you look closely at the goal state above, then you see the tactics @{ML repeat_xmp_tac} |
1159 and @{ML repeat_xmp_tac'} are not yet quite what we are after: the problem is |
1156 and @{ML repeat_xmp_tac'} are not yet quite what we are after: the problem is |
1160 that goals 2 and 3 are not analysed. This is because the tactic |
1157 that goals 2 and 3 are not analysed. This is because the tactic |
1161 is applied repeatedly only to the first subgoal. To analyse also all |
1158 is applied repeatedly only to the first subgoal. To analyse also all |
1162 resulting subgoals, you can use the tactic combinator @{ML_ind REPEAT_ALL_NEW in Tactical}. |
1159 resulting subgoals, you can use the tactic combinator @{ML_ind REPEAT_ALL_NEW in Tactical}. |
1163 Supposing the tactic |
1160 Supposing the tactic |
1164 *} |
1161 \<close> |
1165 |
1162 |
1166 ML %grayML{*fun repeat_all_new_xmp_tac ctxt = REPEAT_ALL_NEW (CHANGED o select_tac' ctxt)*} |
1163 ML %grayML\<open>fun repeat_all_new_xmp_tac ctxt = REPEAT_ALL_NEW (CHANGED o select_tac' ctxt)\<close> |
1167 |
1164 |
1168 text {* |
1165 text \<open> |
1169 you can see that the following goal |
1166 you can see that the following goal |
1170 *} |
1167 \<close> |
1171 |
1168 |
1172 lemma |
1169 lemma |
1173 shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)" |
1170 shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)" |
1174 apply(tactic {* repeat_all_new_xmp_tac @{context} 1 *}) |
1171 apply(tactic \<open>repeat_all_new_xmp_tac @{context} 1\<close>) |
1175 txt{* \begin{minipage}{\textwidth} |
1172 txt\<open>\begin{minipage}{\textwidth} |
1176 @{subgoals [display]} |
1173 @{subgoals [display]} |
1177 \end{minipage} *} |
1174 \end{minipage}\<close> |
1178 (*<*)oops(*>*) |
1175 (*<*)oops(*>*) |
1179 |
1176 |
1180 text {* |
1177 text \<open> |
1181 is completely analysed according to the theorems we chose to |
1178 is completely analysed according to the theorems we chose to |
1182 include in @{ML select_tac'}. |
1179 include in @{ML select_tac'}. |
1183 |
1180 |
1184 Recall that tactics produce a lazy sequence of successor goal states. These |
1181 Recall that tactics produce a lazy sequence of successor goal states. These |
1185 states can be explored using the command \isacommand{back}. For example |
1182 states can be explored using the command \isacommand{back}. For example |
1186 |
1183 |
1187 *} |
1184 \<close> |
1188 |
1185 |
1189 lemma |
1186 lemma |
1190 shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" |
1187 shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" |
1191 apply(tactic {* eresolve_tac @{context} [@{thm disjE}] 1 *}) |
1188 apply(tactic \<open>eresolve_tac @{context} [@{thm disjE}] 1\<close>) |
1192 txt{* applies the rule to the first assumption yielding the goal state: |
1189 txt\<open>applies the rule to the first assumption yielding the goal state: |
1193 \begin{isabelle} |
1190 \begin{isabelle} |
1194 @{subgoals [display]} |
1191 @{subgoals [display]} |
1195 \end{isabelle} |
1192 \end{isabelle} |
1196 |
1193 |
1197 After typing |
1194 After typing |
1198 *} |
1195 \<close> |
1199 (*<*) |
1196 (*<*) |
1200 oops |
1197 oops |
1201 lemma |
1198 lemma |
1202 shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" |
1199 shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" |
1203 apply(tactic {* eresolve_tac @{context} [@{thm disjE}] 1 *}) |
1200 apply(tactic \<open>eresolve_tac @{context} [@{thm disjE}] 1\<close>) |
1204 (*>*) |
1201 (*>*) |
1205 back |
1202 back |
1206 txt{* the rule now applies to the second assumption. |
1203 txt\<open>the rule now applies to the second assumption. |
1207 \begin{isabelle} |
1204 \begin{isabelle} |
1208 @{subgoals [display]} |
1205 @{subgoals [display]} |
1209 \end{isabelle} *} |
1206 \end{isabelle}\<close> |
1210 (*<*)oops(*>*) |
1207 (*<*)oops(*>*) |
1211 |
1208 |
1212 text {* |
1209 text \<open> |
1213 Sometimes this leads to confusing behaviour of tactics and also has |
1210 Sometimes this leads to confusing behaviour of tactics and also has |
1214 the potential to explode the search space for tactics. |
1211 the potential to explode the search space for tactics. |
1215 These problems can be avoided by prefixing the tactic with the tactic |
1212 These problems can be avoided by prefixing the tactic with the tactic |
1216 combinator @{ML_ind DETERM in Tactical}. |
1213 combinator @{ML_ind DETERM in Tactical}. |
1217 *} |
1214 \<close> |
1218 |
1215 |
1219 lemma |
1216 lemma |
1220 shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" |
1217 shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R" |
1221 apply(tactic {* DETERM (eresolve_tac @{context} [@{thm disjE}] 1) *}) |
1218 apply(tactic \<open>DETERM (eresolve_tac @{context} [@{thm disjE}] 1)\<close>) |
1222 txt {*\begin{minipage}{\textwidth} |
1219 txt \<open>\begin{minipage}{\textwidth} |
1223 @{subgoals [display]} |
1220 @{subgoals [display]} |
1224 \end{minipage} *} |
1221 \end{minipage}\<close> |
1225 (*<*)oops(*>*) |
1222 (*<*)oops(*>*) |
1226 text {* |
1223 text \<open> |
1227 This combinator will prune the search space to just the first successful application. |
1224 This combinator will prune the search space to just the first successful application. |
1228 Attempting to apply \isacommand{back} in this goal states gives the |
1225 Attempting to apply \isacommand{back} in this goal states gives the |
1229 error message: |
1226 error message: |
1230 |
1227 |
1231 \begin{isabelle} |
1228 \begin{isabelle} |
1232 @{text "*** back: no alternatives"}\\ |
1229 \<open>*** back: no alternatives\<close>\\ |
1233 @{text "*** At command \"back\"."} |
1230 \<open>*** At command "back".\<close> |
1234 \end{isabelle} |
1231 \end{isabelle} |
1235 |
1232 |
1236 |
1233 |
1237 \footnote{\bf FIXME: say something about @{ML_ind COND in Tactical} and COND'} |
1234 \footnote{\bf FIXME: say something about @{ML_ind COND in Tactical} and COND'} |
1238 \footnote{\bf FIXME: PARALLEL-CHOICE PARALLEL-GOALS} |
1235 \footnote{\bf FIXME: PARALLEL-CHOICE PARALLEL-GOALS} |
1239 |
1236 |
1240 \begin{readmore} |
1237 \begin{readmore} |
1241 Most tactic combinators described in this section are defined in @{ML_file "Pure/tactical.ML"}. |
1238 Most tactic combinators described in this section are defined in @{ML_file "Pure/tactical.ML"}. |
1242 Some combinators for the purpose of proof search are implemented in @{ML_file "Pure/search.ML"}. |
1239 Some combinators for the purpose of proof search are implemented in @{ML_file "Pure/search.ML"}. |
1243 \end{readmore} |
1240 \end{readmore} |
1244 *} |
1241 \<close> |
1245 |
1242 |
1246 text {* |
1243 text \<open> |
1247 \begin{exercise}\label{ex:dyckhoff} |
1244 \begin{exercise}\label{ex:dyckhoff} |
1248 Dyckhoff presents in \cite{Dyckhoff92} inference rules of a sequent |
1245 Dyckhoff presents in \cite{Dyckhoff92} inference rules of a sequent |
1249 calculus, called G4ip, for intuitionistic propositional logic. The |
1246 calculus, called G4ip, for intuitionistic propositional logic. The |
1250 interesting feature of this calculus is that no contraction rule is needed |
1247 interesting feature of this calculus is that no contraction rule is needed |
1251 in order to be complete. As a result the rules can be applied exhaustively, which |
1248 in order to be complete. As a result the rules can be applied exhaustively, which |
1758 |
1755 |
1759 (FIXME: what are @{ML mksimps_pairs}; used in Nominal.thy) |
1756 (FIXME: what are @{ML mksimps_pairs}; used in Nominal.thy) |
1760 |
1757 |
1761 (FIXME: explain @{ML simplify} and @{ML "Simplifier.rewrite_rule"} etc.) |
1758 (FIXME: explain @{ML simplify} and @{ML "Simplifier.rewrite_rule"} etc.) |
1762 |
1759 |
1763 *} |
1760 \<close> |
1764 |
1761 |
1765 section {* Simprocs *} |
1762 section \<open>Simprocs\<close> |
1766 |
1763 |
1767 text {* |
1764 text \<open> |
1768 In Isabelle you can also implement custom simplification procedures, called |
1765 In Isabelle you can also implement custom simplification procedures, called |
1769 \emph{simprocs}. Simprocs can be triggered by the simplifier on a specified |
1766 \emph{simprocs}. Simprocs can be triggered by the simplifier on a specified |
1770 term-pattern and rewrite a term according to a theorem. They are useful in |
1767 term-pattern and rewrite a term according to a theorem. They are useful in |
1771 cases where a rewriting rule must be produced on ``demand'' or when |
1768 cases where a rewriting rule must be produced on ``demand'' or when |
1772 rewriting by simplification is too unpredictable and potentially loops. |
1769 rewriting by simplification is too unpredictable and potentially loops. |
1773 |
1770 |
1774 To see how simprocs work, let us first write a simproc that just prints out |
1771 To see how simprocs work, let us first write a simproc that just prints out |
1775 the pattern which triggers it and otherwise does nothing. For this |
1772 the pattern which triggers it and otherwise does nothing. For this |
1776 you can use the function: |
1773 you can use the function: |
1777 *} |
1774 \<close> |
1778 |
1775 |
1779 ML %linenosgray{*fun fail_simproc ctxt redex = |
1776 ML %linenosgray\<open>fun fail_simproc ctxt redex = |
1780 let |
1777 let |
1781 val _ = pwriteln (Pretty.block |
1778 val _ = pwriteln (Pretty.block |
1782 [Pretty.str "The redex: ", pretty_cterm ctxt redex]) |
1779 [Pretty.str "The redex: ", pretty_cterm ctxt redex]) |
1783 in |
1780 in |
1784 NONE |
1781 NONE |
1785 end*} |
1782 end\<close> |
1786 |
1783 |
1787 text {* |
1784 text \<open> |
1788 This function takes a simpset and a redex (a @{ML_type cterm}) as |
1785 This function takes a simpset and a redex (a @{ML_type cterm}) as |
1789 arguments. In Lines 3 and~4, we first extract the context from the given |
1786 arguments. In Lines 3 and~4, we first extract the context from the given |
1790 simpset and then print out a message containing the redex. The function |
1787 simpset and then print out a message containing the redex. The function |
1791 returns @{ML NONE} (standing for an optional @{ML_type thm}) since at the |
1788 returns @{ML NONE} (standing for an optional @{ML_type thm}) since at the |
1792 moment we are \emph{not} interested in actually rewriting anything. We want |
1789 moment we are \emph{not} interested in actually rewriting anything. We want |
1793 that the simproc is triggered by the pattern @{term "Suc n"}. This can be |
1790 that the simproc is triggered by the pattern @{term "Suc n"}. This can be |
1794 done by adding the simproc to the current simpset as follows |
1791 done by adding the simproc to the current simpset as follows |
1795 *} |
1792 \<close> |
1796 |
1793 |
1797 simproc_setup %gray fail ("Suc n") = {* K fail_simproc *} |
1794 simproc_setup %gray fail ("Suc n") = \<open>K fail_simproc\<close> |
1798 |
1795 |
1799 text {* |
1796 text \<open> |
1800 where the second argument specifies the pattern and the right-hand side |
1797 where the second argument specifies the pattern and the right-hand side |
1801 contains the code of the simproc (we have to use @{ML K} since we are ignoring |
1798 contains the code of the simproc (we have to use @{ML K} since we are ignoring |
1802 an argument about morphisms. |
1799 an argument about morphisms. |
1803 After this, the simplifier is aware of the simproc and you can test whether |
1800 After this, the simplifier is aware of the simproc and you can test whether |
1804 it fires on the lemma: |
1801 it fires on the lemma: |
1805 *} |
1802 \<close> |
1806 |
1803 |
1807 lemma |
1804 lemma |
1808 shows "Suc 0 = 1" |
1805 shows "Suc 0 = 1" |
1809 apply(simp) |
1806 apply(simp) |
1810 txt{* |
1807 txt\<open> |
1811 \begin{minipage}{\textwidth} |
1808 \begin{minipage}{\textwidth} |
1812 \small@{text "> The redex: Suc 0"}\\ |
1809 \small\<open>> The redex: Suc 0\<close>\\ |
1813 @{text "> The redex: Suc 0"}\\ |
1810 \<open>> The redex: Suc 0\<close>\\ |
1814 \end{minipage}*}(*<*)oops(*>*) |
1811 \end{minipage}\<close>(*<*)oops(*>*) |
1815 |
1812 |
1816 text {* |
1813 text \<open> |
1817 This will print out the message twice: once for the left-hand side and |
1814 This will print out the message twice: once for the left-hand side and |
1818 once for the right-hand side. The reason is that during simplification the |
1815 once for the right-hand side. The reason is that during simplification the |
1819 simplifier will at some point reduce the term @{term "1::nat"} to @{term "Suc |
1816 simplifier will at some point reduce the term @{term "1::nat"} to @{term "Suc |
1820 0"}, and then the simproc ``fires'' also on that term. |
1817 0"}, and then the simproc ``fires'' also on that term. |
1821 |
1818 |
1822 We can add or delete the simproc from the current simpset by the usual |
1819 We can add or delete the simproc from the current simpset by the usual |
1823 \isacommand{declare}-statement. For example the simproc will be deleted |
1820 \isacommand{declare}-statement. For example the simproc will be deleted |
1824 with the declaration |
1821 with the declaration |
1825 *} |
1822 \<close> |
1826 |
1823 |
1827 declare [[simproc del: fail]] |
1824 declare [[simproc del: fail]] |
1828 |
1825 |
1829 text {* |
1826 text \<open> |
1830 If you want to see what happens with just \emph{this} simproc, without any |
1827 If you want to see what happens with just \emph{this} simproc, without any |
1831 interference from other rewrite rules, you can call @{text fail} |
1828 interference from other rewrite rules, you can call \<open>fail\<close> |
1832 as follows: |
1829 as follows: |
1833 *} |
1830 \<close> |
1834 |
1831 |
1835 lemma |
1832 lemma |
1836 shows "Suc 0 = 1" |
1833 shows "Suc 0 = 1" |
1837 apply(tactic {* simp_tac (put_simpset |
1834 apply(tactic \<open>simp_tac (put_simpset |
1838 HOL_basic_ss @{context} addsimprocs [@{simproc fail}]) 1*}) |
1835 HOL_basic_ss @{context} addsimprocs [@{simproc fail}]) 1\<close>) |
1839 (*<*)oops(*>*) |
1836 (*<*)oops(*>*) |
1840 |
1837 |
1841 text {* |
1838 text \<open> |
1842 Now the message shows up only once since the term @{term "1::nat"} is |
1839 Now the message shows up only once since the term @{term "1::nat"} is |
1843 left unchanged. |
1840 left unchanged. |
1844 |
1841 |
1845 Setting up a simproc using the command \isacommand{simproc\_setup} will |
1842 Setting up a simproc using the command \isacommand{simproc\_setup} will |
1846 always add automatically the simproc to the current simpset. If you do not |
1843 always add automatically the simproc to the current simpset. If you do not |
1847 want this, then you have to use a slightly different method for setting |
1844 want this, then you have to use a slightly different method for setting |
1848 up the simproc. First the function @{ML fail_simproc} needs to be modified |
1845 up the simproc. First the function @{ML fail_simproc} needs to be modified |
1849 to |
1846 to |
1850 *} |
1847 \<close> |
1851 |
1848 |
1852 ML %grayML{*fun fail_simproc' _ ctxt redex = |
1849 ML %grayML\<open>fun fail_simproc' _ ctxt redex = |
1853 let |
1850 let |
1854 val _ = pwriteln (Pretty.block |
1851 val _ = pwriteln (Pretty.block |
1855 [Pretty.str "The redex:", pretty_cterm ctxt redex]) |
1852 [Pretty.str "The redex:", pretty_cterm ctxt redex]) |
1856 in |
1853 in |
1857 NONE |
1854 NONE |
1858 end*} |
1855 end\<close> |
1859 |
1856 |
1860 text {* |
1857 text \<open> |
1861 Here the redex is given as a @{ML_type term}, instead of a @{ML_type cterm} |
1858 Here the redex is given as a @{ML_type term}, instead of a @{ML_type cterm} |
1862 (therefore we printing it out using the function @{ML pretty_term in Pretty}). |
1859 (therefore we printing it out using the function @{ML pretty_term in Pretty}). |
1863 We can turn this function into a proper simproc using the function |
1860 We can turn this function into a proper simproc using the function |
1864 @{ML Simplifier.make_simproc}: |
1861 @{ML Simplifier.make_simproc}: |
1865 *} |
1862 \<close> |
1866 |
1863 |
1867 ML %grayML{*val fail' = |
1864 ML %grayML\<open>val fail' = |
1868 Simplifier.make_simproc @{context} "fail_simproc'" |
1865 Simplifier.make_simproc @{context} "fail_simproc'" |
1869 {lhss = [@{term "Suc n"}], proc = fail_simproc'}*} |
1866 {lhss = [@{term "Suc n"}], proc = fail_simproc'}\<close> |
1870 |
1867 |
1871 text {* |
1868 text \<open> |
1872 Here the pattern is given as @{ML_type term} (instead of @{ML_type cterm}). |
1869 Here the pattern is given as @{ML_type term} (instead of @{ML_type cterm}). |
1873 The function also takes a list of patterns that can trigger the simproc. |
1870 The function also takes a list of patterns that can trigger the simproc. |
1874 Now the simproc is set up and can be explicitly added using |
1871 Now the simproc is set up and can be explicitly added using |
1875 @{ML_ind addsimprocs in Raw_Simplifier} to a simpset whenever |
1872 @{ML_ind addsimprocs in Raw_Simplifier} to a simpset whenever |
1876 needed. |
1873 needed. |
1877 |
1874 |
1878 Simprocs are applied from inside to outside and from left to right. You can |
1875 Simprocs are applied from inside to outside and from left to right. You can |
1879 see this in the proof |
1876 see this in the proof |
1880 *} |
1877 \<close> |
1881 |
1878 |
1882 lemma |
1879 lemma |
1883 shows "Suc (Suc 0) = (Suc 1)" |
1880 shows "Suc (Suc 0) = (Suc 1)" |
1884 apply(tactic {* simp_tac ((put_simpset HOL_basic_ss @{context}) addsimprocs [fail']) 1*}) |
1881 apply(tactic \<open>simp_tac ((put_simpset HOL_basic_ss @{context}) addsimprocs [fail']) 1\<close>) |
1885 (*<*)oops(*>*) |
1882 (*<*)oops(*>*) |
1886 |
1883 |
1887 text {* |
1884 text \<open> |
1888 The simproc @{ML fail'} prints out the sequence |
1885 The simproc @{ML fail'} prints out the sequence |
1889 |
1886 |
1890 @{text [display] |
1887 @{text [display] |
1891 "> Suc 0 |
1888 "> Suc 0 |
1892 > Suc (Suc 0) |
1889 > Suc (Suc 0) |
1893 > Suc 1"} |
1890 > Suc 1"} |
1894 |
1891 |
1895 To see how a simproc applies a theorem, let us implement a simproc that |
1892 To see how a simproc applies a theorem, let us implement a simproc that |
1896 rewrites terms according to the equation: |
1893 rewrites terms according to the equation: |
1897 *} |
1894 \<close> |
1898 |
1895 |
1899 lemma plus_one: |
1896 lemma plus_one: |
1900 shows "Suc n \<equiv> n + 1" by simp |
1897 shows "Suc n \<equiv> n + 1" by simp |
1901 |
1898 |
1902 text {* |
1899 text \<open> |
1903 Simprocs expect that the given equation is a meta-equation, however the |
1900 Simprocs expect that the given equation is a meta-equation, however the |
1904 equation can contain preconditions (the simproc then will only fire if the |
1901 equation can contain preconditions (the simproc then will only fire if the |
1905 preconditions can be solved). To see that one has relatively precise control over |
1902 preconditions can be solved). To see that one has relatively precise control over |
1906 the rewriting with simprocs, let us further assume we want that the simproc |
1903 the rewriting with simprocs, let us further assume we want that the simproc |
1907 only rewrites terms ``greater'' than @{term "Suc 0"}. For this we can write |
1904 only rewrites terms ``greater'' than @{term "Suc 0"}. For this we can write |
1908 *} |
1905 \<close> |
1909 |
1906 |
1910 |
1907 |
1911 ML %grayML{*fun plus_one_simproc _ ctxt redex = |
1908 ML %grayML\<open>fun plus_one_simproc _ ctxt redex = |
1912 case Thm.term_of redex of |
1909 case Thm.term_of redex of |
1913 @{term "Suc 0"} => NONE |
1910 @{term "Suc 0"} => NONE |
1914 | _ => SOME @{thm plus_one}*} |
1911 | _ => SOME @{thm plus_one}\<close> |
1915 |
1912 |
1916 text {* |
1913 text \<open> |
1917 and set up the simproc as follows. |
1914 and set up the simproc as follows. |
1918 *} |
1915 \<close> |
1919 |
1916 |
1920 ML %grayML{*val plus_one = |
1917 ML %grayML\<open>val plus_one = |
1921 Simplifier.make_simproc @{context} "sproc +1" |
1918 Simplifier.make_simproc @{context} "sproc +1" |
1922 {lhss = [@{term "Suc n"}], proc = plus_one_simproc}*} |
1919 {lhss = [@{term "Suc n"}], proc = plus_one_simproc}\<close> |
1923 |
1920 |
1924 text {* |
1921 text \<open> |
1925 Now the simproc is set up so that it is triggered by terms |
1922 Now the simproc is set up so that it is triggered by terms |
1926 of the form @{term "Suc n"}, but inside the simproc we only produce |
1923 of the form @{term "Suc n"}, but inside the simproc we only produce |
1927 a theorem if the term is not @{term "Suc 0"}. The result you can see |
1924 a theorem if the term is not @{term "Suc 0"}. The result you can see |
1928 in the following proof |
1925 in the following proof |
1929 *} |
1926 \<close> |
1930 |
1927 |
1931 lemma |
1928 lemma |
1932 shows "P (Suc (Suc (Suc 0))) (Suc 0)" |
1929 shows "P (Suc (Suc (Suc 0))) (Suc 0)" |
1933 apply(tactic {* simp_tac (put_simpset HOL_basic_ss @{context} addsimprocs [plus_one]) 1*}) |
1930 apply(tactic \<open>simp_tac (put_simpset HOL_basic_ss @{context} addsimprocs [plus_one]) 1\<close>) |
1934 txt{* |
1931 txt\<open> |
1935 where the simproc produces the goal state |
1932 where the simproc produces the goal state |
1936 \begin{isabelle} |
1933 \begin{isabelle} |
1937 @{subgoals[display]} |
1934 @{subgoals[display]} |
1938 \end{isabelle} |
1935 \end{isabelle} |
1939 *} |
1936 \<close> |
1940 (*<*)oops(*>*) |
1937 (*<*)oops(*>*) |
1941 |
1938 |
1942 text {* |
1939 text \<open> |
1943 As usual with rewriting you have to worry about looping: you already have |
1940 As usual with rewriting you have to worry about looping: you already have |
1944 a loop with @{ML plus_one}, if you apply it with the default simpset (because |
1941 a loop with @{ML plus_one}, if you apply it with the default simpset (because |
1945 the default simpset contains a rule which just does the opposite of @{ML plus_one}, |
1942 the default simpset contains a rule which just does the opposite of @{ML plus_one}, |
1946 namely rewriting @{text [quotes] "+ 1"} to a successor). So you have to be careful |
1943 namely rewriting @{text [quotes] "+ 1"} to a successor). So you have to be careful |
1947 in choosing the right simpset to which you add a simproc. |
1944 in choosing the right simpset to which you add a simproc. |
1948 |
1945 |
1949 Next let us implement a simproc that replaces terms of the form @{term "Suc n"} |
1946 Next let us implement a simproc that replaces terms of the form @{term "Suc n"} |
1950 with the number @{text n} increased by one. First we implement a function that |
1947 with the number \<open>n\<close> increased by one. First we implement a function that |
1951 takes a term and produces the corresponding integer value. |
1948 takes a term and produces the corresponding integer value. |
1952 *} |
1949 \<close> |
1953 |
1950 |
1954 ML %grayML{*fun dest_suc_trm ((Const (@{const_name "Suc"}, _)) $ t) = 1 + dest_suc_trm t |
1951 ML %grayML\<open>fun dest_suc_trm ((Const (@{const_name "Suc"}, _)) $ t) = 1 + dest_suc_trm t |
1955 | dest_suc_trm t = snd (HOLogic.dest_number t)*} |
1952 | dest_suc_trm t = snd (HOLogic.dest_number t)\<close> |
1956 |
1953 |
1957 text {* |
1954 text \<open> |
1958 It uses the library function @{ML_ind dest_number in HOLogic} that transforms |
1955 It uses the library function @{ML_ind dest_number in HOLogic} that transforms |
1959 (Isabelle) terms, like @{term "0::nat"}, @{term "1::nat"}, @{term "2::nat"} and so |
1956 (Isabelle) terms, like @{term "0::nat"}, @{term "1::nat"}, @{term "2::nat"} and so |
1960 on, into integer values. This function raises the exception @{ML TERM}, if |
1957 on, into integer values. This function raises the exception @{ML TERM}, if |
1961 the term is not a number. The next function expects a pair consisting of a term |
1958 the term is not a number. The next function expects a pair consisting of a term |
1962 @{text t} (containing @{term Suc}s) and the corresponding integer value @{text n}. |
1959 \<open>t\<close> (containing @{term Suc}s) and the corresponding integer value \<open>n\<close>. |
1963 *} |
1960 \<close> |
1964 |
1961 |
1965 ML %linenosgray{*fun get_thm ctxt (t, n) = |
1962 ML %linenosgray\<open>fun get_thm ctxt (t, n) = |
1966 let |
1963 let |
1967 val num = HOLogic.mk_number @{typ "nat"} n |
1964 val num = HOLogic.mk_number @{typ "nat"} n |
1968 val goal = Logic.mk_equals (t, num) |
1965 val goal = Logic.mk_equals (t, num) |
1969 in |
1966 in |
1970 Goal.prove ctxt [] [] goal (K (Arith_Data.arith_tac ctxt 1)) |
1967 Goal.prove ctxt [] [] goal (K (Arith_Data.arith_tac ctxt 1)) |
1971 end*} |
1968 end\<close> |
1972 |
1969 |
1973 text {* |
1970 text \<open> |
1974 From the integer value it generates the corresponding number term, called |
1971 From the integer value it generates the corresponding number term, called |
1975 @{text num} (Line 3), and then generates the meta-equation @{text "t \<equiv> num"} |
1972 \<open>num\<close> (Line 3), and then generates the meta-equation \<open>t \<equiv> num\<close> |
1976 (Line 4), which it proves by the arithmetic tactic in Line 6. |
1973 (Line 4), which it proves by the arithmetic tactic in Line 6. |
1977 |
1974 |
1978 For our purpose at the moment, proving the meta-equation using @{ML |
1975 For our purpose at the moment, proving the meta-equation using @{ML |
1979 arith_tac in Arith_Data} is fine, but there is also an alternative employing |
1976 arith_tac in Arith_Data} is fine, but there is also an alternative employing |
1980 the simplifier with a special simpset. For the kind of lemmas we |
1977 the simplifier with a special simpset. For the kind of lemmas we |
1981 want to prove here, the simpset @{text "num_ss"} should suffice. |
1978 want to prove here, the simpset \<open>num_ss\<close> should suffice. |
1982 *} |
1979 \<close> |
1983 |
1980 |
1984 ML %grayML{*fun get_thm_alt ctxt (t, n) = |
1981 ML %grayML\<open>fun get_thm_alt ctxt (t, n) = |
1985 let |
1982 let |
1986 val num = HOLogic.mk_number @{typ "nat"} n |
1983 val num = HOLogic.mk_number @{typ "nat"} n |
1987 val goal = Logic.mk_equals (t, num) |
1984 val goal = Logic.mk_equals (t, num) |
1988 val num_ss = put_simpset HOL_ss ctxt addsimps @{thms semiring_norm} |
1985 val num_ss = put_simpset HOL_ss ctxt addsimps @{thms semiring_norm} |
1989 in |
1986 in |
1990 Goal.prove ctxt [] [] goal (K (simp_tac num_ss 1)) |
1987 Goal.prove ctxt [] [] goal (K (simp_tac num_ss 1)) |
1991 end*} |
1988 end\<close> |
1992 |
1989 |
1993 text {* |
1990 text \<open> |
1994 The advantage of @{ML get_thm_alt} is that it leaves very little room for |
1991 The advantage of @{ML get_thm_alt} is that it leaves very little room for |
1995 something to go wrong; in contrast it is much more difficult to predict |
1992 something to go wrong; in contrast it is much more difficult to predict |
1996 what happens with @{ML arith_tac in Arith_Data}, especially in more complicated |
1993 what happens with @{ML arith_tac in Arith_Data}, especially in more complicated |
1997 circumstances. The disadvantage of @{ML get_thm_alt} is to find a simpset |
1994 circumstances. The disadvantage of @{ML get_thm_alt} is to find a simpset |
1998 that is sufficiently powerful to solve every instance of the lemmas |
1995 that is sufficiently powerful to solve every instance of the lemmas |
2412 \end{itemize} |
2409 \end{itemize} |
2413 |
2410 |
2414 Assume we want to apply @{ML true_conj1_conv} only in the conclusion |
2411 Assume we want to apply @{ML true_conj1_conv} only in the conclusion |
2415 of the goal, and @{ML if_true1_conv} should only apply to the premises. |
2412 of the goal, and @{ML if_true1_conv} should only apply to the premises. |
2416 Here is a tactic doing exactly that: |
2413 Here is a tactic doing exactly that: |
2417 *} |
2414 \<close> |
2418 |
2415 |
2419 ML %grayML{*fun true1_tac ctxt = |
2416 ML %grayML\<open>fun true1_tac ctxt = |
2420 CONVERSION |
2417 CONVERSION |
2421 (Conv.params_conv ~1 (fn ctxt => |
2418 (Conv.params_conv ~1 (fn ctxt => |
2422 (Conv.prems_conv ~1 (if_true1_conv ctxt) then_conv |
2419 (Conv.prems_conv ~1 (if_true1_conv ctxt) then_conv |
2423 Conv.concl_conv ~1 (true_conj1_conv ctxt))) ctxt)*} |
2420 Conv.concl_conv ~1 (true_conj1_conv ctxt))) ctxt)\<close> |
2424 |
2421 |
2425 text {* |
2422 text \<open> |
2426 We call the conversions with the argument @{ML "~1"}. By this we |
2423 We call the conversions with the argument @{ML "~1"}. By this we |
2427 analyse all parameters, all premises and the conclusion of a goal |
2424 analyse all parameters, all premises and the conclusion of a goal |
2428 state. If we call @{ML concl_conv in Conv} with |
2425 state. If we call @{ML concl_conv in Conv} with |
2429 a non-negative number, say @{text n}, then this conversions will |
2426 a non-negative number, say \<open>n\<close>, then this conversions will |
2430 skip the first @{text n} premises and applies the conversion to the |
2427 skip the first \<open>n\<close> premises and applies the conversion to the |
2431 ``rest'' (similar for parameters and conclusions). To test the |
2428 ``rest'' (similar for parameters and conclusions). To test the |
2432 tactic, consider the proof |
2429 tactic, consider the proof |
2433 *} |
2430 \<close> |
2434 |
2431 |
2435 lemma |
2432 lemma |
2436 "if True \<and> P then P else True \<and> False \<Longrightarrow> |
2433 "if True \<and> P then P else True \<and> False \<Longrightarrow> |
2437 (if True \<and> Q then True \<and> Q else P) \<longrightarrow> True \<and> (True \<and> Q)" |
2434 (if True \<and> Q then True \<and> Q else P) \<longrightarrow> True \<and> (True \<and> Q)" |
2438 apply(tactic {* true1_tac @{context} 1 *}) |
2435 apply(tactic \<open>true1_tac @{context} 1\<close>) |
2439 txt {* where the tactic yields the goal state |
2436 txt \<open>where the tactic yields the goal state |
2440 \begin{isabelle} |
2437 \begin{isabelle} |
2441 @{subgoals [display]} |
2438 @{subgoals [display]} |
2442 \end{isabelle}*} |
2439 \end{isabelle}\<close> |
2443 (*<*)oops(*>*) |
2440 (*<*)oops(*>*) |
2444 |
2441 |
2445 text {* |
2442 text \<open> |
2446 As you can see, the premises are rewritten according to @{ML if_true1_conv}, while |
2443 As you can see, the premises are rewritten according to @{ML if_true1_conv}, while |
2447 the conclusion according to @{ML true_conj1_conv}. If we only have one |
2444 the conclusion according to @{ML true_conj1_conv}. If we only have one |
2448 conversion that should be uniformly applied to the whole goal state, we |
2445 conversion that should be uniformly applied to the whole goal state, we |
2449 can simplify @{ML true1_tac} using @{ML_ind top_conv in Conv}. |
2446 can simplify @{ML true1_tac} using @{ML_ind top_conv in Conv}. |
2450 |
2447 |
2451 Conversions are also be helpful for constructing meta-equality theorems. |
2448 Conversions are also be helpful for constructing meta-equality theorems. |
2452 Such theorems are often definitions. As an example consider the following |
2449 Such theorems are often definitions. As an example consider the following |
2453 two ways of defining the identity function in Isabelle. |
2450 two ways of defining the identity function in Isabelle. |
2454 *} |
2451 \<close> |
2455 |
2452 |
2456 definition id1::"'a \<Rightarrow> 'a" |
2453 definition id1::"'a \<Rightarrow> 'a" |
2457 where "id1 x \<equiv> x" |
2454 where "id1 x \<equiv> x" |
2458 |
2455 |
2459 definition id2::"'a \<Rightarrow> 'a" |
2456 definition id2::"'a \<Rightarrow> 'a" |
2460 where "id2 \<equiv> \<lambda>x. x" |
2457 where "id2 \<equiv> \<lambda>x. x" |
2461 |
2458 |
2462 text {* |
2459 text \<open> |
2463 Although both definitions define the same function, the theorems @{thm |
2460 Although both definitions define the same function, the theorems @{thm |
2464 [source] id1_def} and @{thm [source] id2_def} are different meta-equations. However it is |
2461 [source] id1_def} and @{thm [source] id2_def} are different meta-equations. However it is |
2465 easy to transform one into the other. The function @{ML_ind abs_def |
2462 easy to transform one into the other. The function @{ML_ind abs_def |
2466 in Drule} from the structure @{ML_struct Drule} can automatically transform |
2463 in Drule} from the structure @{ML_struct Drule} can automatically transform |
2467 theorem @{thm [source] id1_def} |
2464 theorem @{thm [source] id1_def} |
2492 fun get_conv [] = Conv.rewr_conv def |
2489 fun get_conv [] = Conv.rewr_conv def |
2493 | get_conv (_::xs) = Conv.fun_conv (get_conv xs) |
2490 | get_conv (_::xs) = Conv.fun_conv (get_conv xs) |
2494 in |
2491 in |
2495 get_conv xs new_lhs |> |
2492 get_conv xs new_lhs |> |
2496 singleton (Proof_Context.export ctxt' ctxt) |
2493 singleton (Proof_Context.export ctxt' ctxt) |
2497 end*} |
2494 end\<close> |
2498 |
2495 |
2499 text {* |
2496 text \<open> |
2500 In Line 3 we destruct the meta-equality into the @{ML_type cterm}s |
2497 In Line 3 we destruct the meta-equality into the @{ML_type cterm}s |
2501 corresponding to the left-hand and right-hand side of the meta-equality. The |
2498 corresponding to the left-hand and right-hand side of the meta-equality. The |
2502 assumption in @{ML unabs_def} is that the right-hand side is an |
2499 assumption in @{ML unabs_def} is that the right-hand side is an |
2503 abstraction. We compute the possibly empty list of abstracted variables in |
2500 abstraction. We compute the possibly empty list of abstracted variables in |
2504 Line 4 using the function @{ML_ind strip_abs_vars in Term}. For this we have to |
2501 Line 4 using the function @{ML_ind strip_abs_vars in Term}. For this we have to |
2505 first transform the @{ML_type cterm} into a @{ML_type term}. In Line 5 we |
2502 first transform the @{ML_type cterm} into a @{ML_type term}. In Line 5 we |
2506 import these variables into the context @{text "ctxt'"}, in order to |
2503 import these variables into the context \<open>ctxt'\<close>, in order to |
2507 export them again in Line 15. In Line 8 we certify the list of variables, |
2504 export them again in Line 15. In Line 8 we certify the list of variables, |
2508 which in turn we apply to the @{ML_text lhs} of the definition using the |
2505 which in turn we apply to the @{ML_text lhs} of the definition using the |
2509 function @{ML_ind list_comb in Drule}. In Line 11 and 12 we generate the |
2506 function @{ML_ind list_comb in Drule}. In Line 11 and 12 we generate the |
2510 conversion according to the length of the list of (abstracted) variables. If |
2507 conversion according to the length of the list of (abstracted) variables. If |
2511 there are none, then we just return the conversion corresponding to the |
2508 there are none, then we just return the conversion corresponding to the |
2512 original definition. If there are variables, then we have to prefix this |
2509 original definition. If there are variables, then we have to prefix this |
2513 conversion with @{ML_ind fun_conv in Conv}. Now in Line 14 we only have to |
2510 conversion with @{ML_ind fun_conv in Conv}. Now in Line 14 we only have to |
2514 apply the new left-hand side to the generated conversion and obtain the the |
2511 apply the new left-hand side to the generated conversion and obtain the the |
2515 theorem we want to construct. As mentioned above, in Line 15 we only have to |
2512 theorem we want to construct. As mentioned above, in Line 15 we only have to |
2516 export the context @{text "ctxt'"} back to @{text "ctxt"} in order to |
2513 export the context \<open>ctxt'\<close> back to \<open>ctxt\<close> in order to |
2517 produce meta-variables for the theorem. An example is as follows. |
2514 produce meta-variables for the theorem. An example is as follows. |
2518 |
2515 |
2519 @{ML_response_fake [display, gray] |
2516 @{ML_response_fake [display, gray] |
2520 "unabs_def @{context} @{thm id2_def}" |
2517 "unabs_def @{context} @{thm id2_def}" |
2521 "id2 ?x1 \<equiv> ?x1"} |
2518 "id2 ?x1 \<equiv> ?x1"} |
2522 *} |
2519 \<close> |
2523 |
2520 |
2524 text {* |
2521 text \<open> |
2525 To sum up this section, conversions are more general than the simplifier |
2522 To sum up this section, conversions are more general than the simplifier |
2526 or simprocs, but you have to do more work yourself. Also conversions are |
2523 or simprocs, but you have to do more work yourself. Also conversions are |
2527 often much less efficient than the simplifier. The advantage of conversions, |
2524 often much less efficient than the simplifier. The advantage of conversions, |
2528 however, is that they provide much less room for non-termination. |
2525 however, is that they provide much less room for non-termination. |
2529 |
2526 |