93
+ − 1
theory Tactical
441
+ − 2
imports Base First_Steps
93
+ − 3
begin
+ − 4
565
+ − 5
chapter \<open>Tactical Reasoning\label{chp:tactical}\<close>
+ − 6
+ − 7
text \<open>
506
+ − 8
\begin{flushright}
+ − 9
{\em
+ − 10
``The first thing I would say is that when you write a program, think of\\
+ − 11
it primarily as a work of literature. You're trying to write something\\
+ − 12
that human beings are going to read. Don't think of it primarily as\\
+ − 13
something a computer is going to follow. The more effective you are\\
+ − 14
at making your program readable, the more effective it's going to\\
+ − 15
be: You'll understand it today, you'll understand it next week, and\\
+ − 16
your successors who are going to maintain and modify it will\\
+ − 17
understand it.''}\\[1ex]
+ − 18
Donald E.~Knuth, from an interview in Dr.~Dobb's Journal, 1996.
+ − 19
\end{flushright}
+ − 20
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 21
One of the main reason for descending to the ML-level of Isabelle is to
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 22
implement automatic proof procedures. Such proof procedures can
362
+ − 23
considerably lessen the burden of manual reasoning. They are centred around
+ − 24
the idea of refining a goal state using tactics. This is similar to the
+ − 25
\isacommand{apply}-style reasoning at the user-level, where goals are
+ − 26
modified in a sequence of proof steps until all of them are discharged.
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 27
In this chapter we explain how to implement simple tactics and how to combine them using
363
+ − 28
tactic combinators. We also describe the simplifier, simprocs and conversions.
565
+ − 29
\<close>
+ − 30
+ − 31
section \<open>Basics of Reasoning with Tactics\label{sec:basictactics}\<close>
+ − 32
+ − 33
text \<open>
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 34
To see how tactics work, let us first transcribe a simple \isacommand{apply}-style proof
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 35
into ML. Suppose the following proof.
565
+ − 36
\<close>
93
+ − 37
362
+ − 38
lemma disj_swap:
+ − 39
shows "P \<or> Q \<Longrightarrow> Q \<or> P"
93
+ − 40
apply(erule disjE)
+ − 41
apply(rule disjI2)
+ − 42
apply(assumption)
+ − 43
apply(rule disjI1)
+ − 44
apply(assumption)
+ − 45
done
+ − 46
565
+ − 47
text \<open>
93
+ − 48
This proof translates to the following ML-code.
+ − 49
+ − 50
@{ML_response_fake [display,gray]
+ − 51
"let
+ − 52
val ctxt = @{context}
+ − 53
val goal = @{prop \"P \<or> Q \<Longrightarrow> Q \<or> P\"}
+ − 54
in
+ − 55
Goal.prove ctxt [\"P\", \"Q\"] [] goal
562
+ − 56
(fn _ => eresolve_tac @{context} [@{thm disjE}] 1
+ − 57
THEN resolve_tac @{context} [@{thm disjI2}] 1
+ − 58
THEN assume_tac @{context} 1
+ − 59
THEN resolve_tac @{context} [@{thm disjI1}] 1
+ − 60
THEN assume_tac @{context} 1)
93
+ − 61
end" "?P \<or> ?Q \<Longrightarrow> ?Q \<or> ?P"}
+ − 62
362
+ − 63
To start the proof, the function @{ML_ind prove in Goal} sets up a goal
+ − 64
state for proving the goal @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. We can give this
+ − 65
function some assumptions in the third argument (there are no assumption in
+ − 66
the proof at hand). The second argument stands for a list of variables
+ − 67
(given as strings). This list contains the variables that will be turned
565
+ − 68
into schematic variables once the goal is proved (in our case \<open>P\<close> and
+ − 69
\<open>Q\<close>). The last argument is the tactic that proves the goal. This
362
+ − 70
tactic can make use of the local assumptions (there are none in this
562
+ − 71
example). The tactics @{ML_ind eresolve_tac in Tactic}, @{ML_ind resolve_tac in Tactic} and
565
+ − 72
@{ML_ind assume_tac in Tactic} in the code above correspond roughly to \<open>erule\<close>, \<open>rule\<close> and \<open>assumption\<close>, respectively. The combinator
363
+ − 73
@{ML THEN} strings the tactics together.
93
+ − 74
562
+ − 75
TBD: Write something about @{ML_ind prove_common in Goal}.
437
+ − 76
93
+ − 77
\begin{readmore}
318
+ − 78
To learn more about the function @{ML_ind prove in Goal} see
+ − 79
\isccite{sec:results} and the file @{ML_file "Pure/goal.ML"}. See @{ML_file
289
+ − 80
"Pure/tactic.ML"} and @{ML_file "Pure/tactical.ML"} for the code of basic
99
+ − 81
tactics and tactic combinators; see also Chapters 3 and 4 in the old
318
+ − 82
Isabelle Reference Manual, and Chapter 3 in the Isabelle/Isar Implementation
+ − 83
Manual.
93
+ − 84
\end{readmore}
+ − 85
+ − 86
436
+ − 87
\index{tactic@ {\tt\slshape{}tactic}}
+ − 88
\index{raw\_tactic@ {\tt\slshape{}raw\_tactic}}
318
+ − 89
During the development of automatic proof procedures, you will often find it
+ − 90
necessary to test a tactic on examples. This can be conveniently done with
565
+ − 91
the command \isacommand{apply}\<open>(tactic \<verbopen> \<dots> \<verbclose>)\<close>.
93
+ − 92
Consider the following sequence of tactics
318
+ − 93
565
+ − 94
\<close>
+ − 95
+ − 96
ML %grayML\<open>fun foo_tac ctxt =
562
+ − 97
(eresolve_tac ctxt [@{thm disjE}] 1
+ − 98
THEN resolve_tac ctxt [@{thm disjI2}] 1
+ − 99
THEN assume_tac ctxt 1
+ − 100
THEN resolve_tac ctxt [@{thm disjI1}] 1
565
+ − 101
THEN assume_tac ctxt 1)\<close>
+ − 102
+ − 103
text \<open>and the Isabelle proof:\<close>
93
+ − 104
362
+ − 105
lemma
+ − 106
shows "P \<or> Q \<Longrightarrow> Q \<or> P"
565
+ − 107
apply(tactic \<open>foo_tac @{context}\<close>)
93
+ − 108
done
+ − 109
565
+ − 110
text \<open>
+ − 111
By using \<open>tactic \<verbopen> \<dots> \<verbclose>\<close> you can call from the
157
+ − 112
user-level of Isabelle the tactic @{ML foo_tac} or
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 113
any other function that returns a tactic. There are some goal transformations
565
+ − 114
that are performed by \<open>tactic\<close>. They can be avoided by using
+ − 115
\<open>raw_tactic\<close> instead.
93
+ − 116
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 117
The tactic @{ML foo_tac} is just a sequence of simple tactics stringed
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 118
together by @{ML THEN}. As can be seen, each simple tactic in @{ML foo_tac}
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 119
has a hard-coded number that stands for the subgoal analysed by the
565
+ − 120
tactic (\<open>1\<close> stands for the first, or top-most, subgoal). This hard-coding
105
+ − 121
of goals is sometimes wanted, but usually it is not. To avoid the explicit numbering,
238
+ − 122
you can write
565
+ − 123
\<close>
+ − 124
+ − 125
ML %grayML\<open>fun foo_tac' ctxt =
562
+ − 126
(eresolve_tac ctxt [@{thm disjE}]
+ − 127
THEN' resolve_tac ctxt [@{thm disjI2}]
+ − 128
THEN' assume_tac ctxt
+ − 129
THEN' resolve_tac ctxt [@{thm disjI1}]
565
+ − 130
THEN' assume_tac ctxt)\<close>text_raw\<open>\label{tac:footacprime}\<close>
+ − 131
+ − 132
text \<open>
363
+ − 133
where @{ML_ind THEN' in Tactical} is used instead of @{ML THEN in
+ − 134
Tactical}. (For most combinators that combine tactics---@{ML THEN} is only
+ − 135
one such combinator---a ``primed'' version exists.) With @{ML foo_tac'} you
+ − 136
can give the number for the subgoal explicitly when the tactic is called. So
+ − 137
in the next proof you can first discharge the second subgoal, and
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 138
then the first.
565
+ − 139
\<close>
93
+ − 140
362
+ − 141
lemma
+ − 142
shows "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1"
+ − 143
and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2"
565
+ − 144
apply(tactic \<open>foo_tac' @{context} 2\<close>)
+ − 145
apply(tactic \<open>foo_tac' @{context} 1\<close>)
93
+ − 146
done
+ − 147
565
+ − 148
text \<open>
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 149
This kind of addressing is more difficult to achieve when the goal is
362
+ − 150
hard-coded inside the tactic.
+ − 151
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 152
The tactics @{ML foo_tac} and @{ML foo_tac'} are specific
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 153
to goals of the form @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. If the goal is not of
362
+ − 154
this form, then these tactics return the error message:\footnote{To be
363
+ − 155
precise, tactics do not produce this error message; the message originates from the
+ − 156
\isacommand{apply} wrapper that calls the tactic.}
362
+ − 157
99
+ − 158
+ − 159
\begin{isabelle}
565
+ − 160
\<open>*** empty result sequence -- proof command failed\<close>\\
+ − 161
\<open>*** At command "apply".\<close>
99
+ − 162
\end{isabelle}
+ − 163
362
+ − 164
This means they failed. The reason for this error message is that tactics
+ − 165
are functions mapping a goal state to a (lazy) sequence of successor
+ − 166
states. Hence the type of a tactic is:
565
+ − 167
\<close>
93
+ − 168
565
+ − 169
ML %grayML\<open>type tactic = thm -> thm Seq.seq\<close>
+ − 170
+ − 171
text \<open>
109
+ − 172
By convention, if a tactic fails, then it should return the empty sequence.
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 173
Therefore, your own tactics should not raise exceptions
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 174
willy-nilly; only in very grave failure situations should a tactic raise the
565
+ − 175
exception \<open>THM\<close>.
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 176
363
+ − 177
The simplest tactics are @{ML_ind no_tac in Tactical} and
+ − 178
@{ML_ind all_tac in Tactical}. The first returns the empty sequence and
+ − 179
is defined as
565
+ − 180
\<close>
+ − 181
+ − 182
ML %grayML\<open>fun no_tac thm = Seq.empty\<close>
+ − 183
+ − 184
text \<open>
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 185
which means @{ML no_tac} always fails. The second returns the given theorem wrapped
173
+ − 186
in a single member sequence; it is defined as
565
+ − 187
\<close>
+ − 188
+ − 189
ML %grayML\<open>fun all_tac thm = Seq.single thm\<close>
+ − 190
+ − 191
text \<open>
109
+ − 192
which means @{ML all_tac} always succeeds, but also does not make any progress
+ − 193
with the proof.
93
+ − 194
109
+ − 195
The lazy list of possible successor goal states shows through at the user-level
99
+ − 196
of Isabelle when using the command \isacommand{back}. For instance in the
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 197
following proof there are two possibilities for how to apply
107
+ − 198
@{ML foo_tac'}: either using the first assumption or the second.
565
+ − 199
\<close>
93
+ − 200
362
+ − 201
lemma
+ − 202
shows "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P"
565
+ − 203
apply(tactic \<open>foo_tac' @{context} 1\<close>)
93
+ − 204
back
+ − 205
done
+ − 206
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 207
565
+ − 208
text \<open>
99
+ − 209
By using \isacommand{back}, we construct the proof that uses the
109
+ − 210
second assumption. While in the proof above, it does not really matter which
+ − 211
assumption is used, in more interesting cases provability might depend
+ − 212
on exploring different possibilities.
99
+ − 213
93
+ − 214
\begin{readmore}
+ − 215
See @{ML_file "Pure/General/seq.ML"} for the implementation of lazy
109
+ − 216
sequences. In day-to-day Isabelle programming, however, one rarely
+ − 217
constructs sequences explicitly, but uses the predefined tactics and
+ − 218
tactic combinators instead.
93
+ − 219
\end{readmore}
+ − 220
104
+ − 221
It might be surprising that tactics, which transform
109
+ − 222
one goal state to the next, are functions from theorems to theorem
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 223
(sequences). The surprise resolves by knowing that every
104
+ − 224
goal state is indeed a theorem. To shed more light on this,
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 225
let us modify the code of @{ML all_tac} to obtain the following
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 226
tactic
565
+ − 227
\<close>
+ − 228
+ − 229
ML %grayML\<open>fun my_print_tac ctxt thm =
132
+ − 230
let
440
+ − 231
val _ = tracing (Pretty.string_of (pretty_thm_no_vars ctxt thm))
132
+ − 232
in
+ − 233
Seq.single thm
565
+ − 234
end\<close>
+ − 235
+ − 236
text \<open>
363
+ − 237
which prints out the given theorem (using the string-function defined in
+ − 238
Section~\ref{sec:printing}) and then behaves like @{ML all_tac}. With this
+ − 239
tactic we are in the position to inspect every goal state in a proof. For
+ − 240
this consider the proof in Figure~\ref{fig:goalstates}: as can be seen,
+ − 241
internally every goal state is an implication of the form
+ − 242
551
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 243
@{text[display] "A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> #C"}
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 245
where @{term C} is the goal to be proved and the @{term "A\<^sub>i"} are
363
+ − 246
the subgoals. So after setting up the lemma, the goal state is always of the
565
+ − 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
363
+ − 248
``protector'' wrapped around it (the wrapper is the outermost constant
565
+ − 249
\<open>Const ("Pure.prop", bool \<Rightarrow> bool)\<close>; in the figure we make it visible
+ − 250
as a \<open>#\<close>). This wrapper prevents that premises of \<open>C\<close> are
363
+ − 251
misinterpreted as open subgoals. While tactics can operate on the subgoals
565
+ − 252
(the \<open>A\<^sub>i\<close> above), they are expected to leave the conclusion
363
+ − 253
@{term C} intact, with the exception of possibly instantiating schematic
+ − 254
variables. This instantiations of schematic variables can be observed
+ − 255
on the user level. Have a look at the following definition and proof.
565
+ − 256
\<close>
+ − 257
+ − 258
text_raw \<open>
109
+ − 259
\begin{figure}[p]
173
+ − 260
\begin{boxedminipage}{\textwidth}
109
+ − 261
\begin{isabelle}
565
+ − 262
\<close>
554
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 263
notation (output) "Pure.prop" ("#_" [1000] 1000)
303
05e6a33edef6
added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 264
362
+ − 265
lemma
+ − 266
shows "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"
565
+ − 267
apply(tactic \<open>my_print_tac @{context}\<close>)
+ − 268
+ − 269
txt\<open>\begin{minipage}{\textwidth}
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 270
@{subgoals [display]}
109
+ − 271
\end{minipage}\medskip
+ − 272
+ − 273
\begin{minipage}{\textwidth}
+ − 274
\small\colorbox{gray!20}{
+ − 275
\begin{tabular}{@ {}l@ {}}
+ − 276
internal goal state:\\
303
05e6a33edef6
added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 277
@{raw_goal_state}
109
+ − 278
\end{tabular}}
+ − 279
\end{minipage}\medskip
565
+ − 280
\<close>
93
+ − 281
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 282
apply(rule conjI)
565
+ − 283
apply(tactic \<open>my_print_tac @{context}\<close>)
+ − 284
+ − 285
txt\<open>\begin{minipage}{\textwidth}
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 286
@{subgoals [display]}
109
+ − 287
\end{minipage}\medskip
+ − 288
+ − 289
\begin{minipage}{\textwidth}
+ − 290
\small\colorbox{gray!20}{
+ − 291
\begin{tabular}{@ {}l@ {}}
+ − 292
internal goal state:\\
303
05e6a33edef6
added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 293
@{raw_goal_state}
109
+ − 294
\end{tabular}}
+ − 295
\end{minipage}\medskip
565
+ − 296
\<close>
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 297
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 298
apply(assumption)
565
+ − 299
apply(tactic \<open>my_print_tac @{context}\<close>)
+ − 300
+ − 301
txt\<open>\begin{minipage}{\textwidth}
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 302
@{subgoals [display]}
109
+ − 303
\end{minipage}\medskip
+ − 304
+ − 305
\begin{minipage}{\textwidth}
+ − 306
\small\colorbox{gray!20}{
+ − 307
\begin{tabular}{@ {}l@ {}}
+ − 308
internal goal state:\\
303
05e6a33edef6
added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 309
@{raw_goal_state}
109
+ − 310
\end{tabular}}
+ − 311
\end{minipage}\medskip
565
+ − 312
\<close>
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 313
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 314
apply(assumption)
565
+ − 315
apply(tactic \<open>my_print_tac @{context}\<close>)
+ − 316
+ − 317
txt\<open>\begin{minipage}{\textwidth}
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 318
@{subgoals [display]}
109
+ − 319
\end{minipage}\medskip
+ − 320
+ − 321
\begin{minipage}{\textwidth}
+ − 322
\small\colorbox{gray!20}{
+ − 323
\begin{tabular}{@ {}l@ {}}
+ − 324
internal goal state:\\
303
05e6a33edef6
added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 325
@{raw_goal_state}
109
+ − 326
\end{tabular}}
+ − 327
\end{minipage}\medskip
565
+ − 328
\<close>
303
05e6a33edef6
added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 329
(*<*)oops(*>*)
565
+ − 330
text_raw \<open>
109
+ − 331
\end{isabelle}
173
+ − 332
\end{boxedminipage}
362
+ − 333
\caption{The figure shows an Isabelle snippet of a proof where each
+ − 334
intermediate goal state is printed by the Isabelle system and by @{ML
+ − 335
my_print_tac}. The latter shows the goal state as represented internally
+ − 336
(highlighted boxes). This tactic shows that every goal state in Isabelle is
565
+ − 337
represented by a theorem: when you start the proof of \mbox{\<open>\<lbrakk>A; B\<rbrakk> \<Longrightarrow>
+ − 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
+ − 339
you finish the proof the theorem is \<open>#(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and>
+ − 340
B)\<close>.\label{fig:goalstates}}
109
+ − 341
\end{figure}
565
+ − 342
\<close>
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 343
358
+ − 344
definition
+ − 345
EQ_TRUE
+ − 346
where
+ − 347
"EQ_TRUE X \<equiv> (X = True)"
+ − 348
562
+ − 349
schematic_goal test:
358
+ − 350
shows "EQ_TRUE ?X"
363
+ − 351
unfolding EQ_TRUE_def
+ − 352
by (rule refl)
358
+ − 353
565
+ − 354
text \<open>
422
+ − 355
By using \isacommand{schematic\_lemma} it is possible to prove lemmas including
565
+ − 356
meta-variables on the user level. However, the proved theorem is not \<open>EQ_TRUE ?X\<close>,
422
+ − 357
as one might expect, but @{thm test}. We can test this with:
358
+ − 358
+ − 359
\begin{isabelle}
+ − 360
\isacommand{thm}~@{thm [source] test}\\
565
+ − 361
\<open>>\<close>~@{thm test}
358
+ − 362
\end{isabelle}
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 363
565
+ − 364
The reason for this result is that the schematic variable \<open>?X\<close>
363
+ − 365
is instantiated inside the proof to be @{term True} and then the
+ − 366
instantiation propagated to the ``outside''.
359
+ − 367
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 368
\begin{readmore}
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 369
For more information about the internals of goals see \isccite{sec:tactical-goals}.
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 370
\end{readmore}
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 371
565
+ − 372
\<close>
+ − 373
+ − 374
+ − 375
section \<open>Simple Tactics\label{sec:simpletacs}\<close>
+ − 376
+ − 377
text \<open>
368
+ − 378
In this section we will introduce more of the simple tactics in Isabelle. The
363
+ − 379
first one is @{ML_ind print_tac in Tactical}, which is quite useful
173
+ − 380
for low-level debugging of tactics. It just prints out a message and the current
+ − 381
goal state. Unlike @{ML my_print_tac} shown earlier, it prints the goal state
+ − 382
as the user would see it. For example, processing the proof
565
+ − 383
\<close>
105
+ − 384
362
+ − 385
lemma
+ − 386
shows "False \<Longrightarrow> True"
565
+ − 387
apply(tactic \<open>print_tac @{context} "foo message"\<close>)
+ − 388
txt\<open>gives:
368
+ − 389
\begin{isabelle}
565
+ − 390
\<open>foo message\<close>\\[3mm]
107
+ − 391
@{prop "False \<Longrightarrow> True"}\\
565
+ − 392
\<open> 1. False \<Longrightarrow> True\<close>\\
368
+ − 393
\end{isabelle}
565
+ − 394
\<close>
105
+ − 395
(*<*)oops(*>*)
+ − 396
565
+ − 397
text \<open>
363
+ − 398
A simple tactic for easy discharge of any proof obligations, even difficult ones, is
+ − 399
@{ML_ind cheat_tac in Skip_Proof} in the structure @{ML_struct Skip_Proof}.
362
+ − 400
This tactic corresponds to the Isabelle command \isacommand{sorry} and is
+ − 401
sometimes useful during the development of tactics.
565
+ − 402
\<close>
192
+ − 403
362
+ − 404
lemma
+ − 405
shows "False" and "Goldbach_conjecture"
565
+ − 406
apply(tactic \<open>Skip_Proof.cheat_tac @{context} 1\<close>)
+ − 407
txt\<open>\begin{minipage}{\textwidth}
192
+ − 408
@{subgoals [display]}
565
+ − 409
\end{minipage}\<close>
192
+ − 410
(*<*)oops(*>*)
+ − 411
565
+ − 412
text \<open>
562
+ − 413
Another simple tactic is the function @{ML_ind assume_tac in Tactic}, which, as shown
362
+ − 414
earlier, corresponds to the assumption method.
565
+ − 415
\<close>
99
+ − 416
362
+ − 417
lemma
+ − 418
shows "P \<Longrightarrow> P"
565
+ − 419
apply(tactic \<open>assume_tac @{context} 1\<close>)
+ − 420
txt\<open>\begin{minipage}{\textwidth}
109
+ − 421
@{subgoals [display]}
565
+ − 422
\end{minipage}\<close>
109
+ − 423
(*<*)oops(*>*)
93
+ − 424
565
+ − 425
text \<open>
562
+ − 426
Similarly, @{ML_ind resolve_tac in Tactic}, @{ML_ind dresolve_tac in Tactic}, @{ML_ind eresolve_tac
565
+ − 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
363
+ − 428
them takes a theorem as argument and attempts to apply it to a goal. Below
+ − 429
are three self-explanatory examples.
565
+ − 430
\<close>
99
+ − 431
362
+ − 432
lemma
+ − 433
shows "P \<and> Q"
565
+ − 434
apply(tactic \<open>resolve_tac @{context} [@{thm conjI}] 1\<close>)
+ − 435
txt\<open>\begin{minipage}{\textwidth}
104
+ − 436
@{subgoals [display]}
565
+ − 437
\end{minipage}\<close>
93
+ − 438
(*<*)oops(*>*)
+ − 439
362
+ − 440
lemma
+ − 441
shows "P \<and> Q \<Longrightarrow> False"
565
+ − 442
apply(tactic \<open>eresolve_tac @{context} [@{thm conjE}] 1\<close>)
+ − 443
txt\<open>\begin{minipage}{\textwidth}
104
+ − 444
@{subgoals [display]}
565
+ − 445
\end{minipage}\<close>
93
+ − 446
(*<*)oops(*>*)
+ − 447
362
+ − 448
lemma
+ − 449
shows "False \<and> True \<Longrightarrow> False"
565
+ − 450
apply(tactic \<open>dresolve_tac @{context} [@{thm conjunct2}] 1\<close>)
+ − 451
txt\<open>\begin{minipage}{\textwidth}
104
+ − 452
@{subgoals [display]}
565
+ − 453
\end{minipage}\<close>
93
+ − 454
(*<*)oops(*>*)
+ − 455
565
+ − 456
text \<open>
562
+ − 457
The function @{ML_ind resolve_tac in Tactic}
363
+ − 458
expects a list of theorems as argument. From this list it will apply the
105
+ − 459
first applicable theorem (later theorems that are also applicable can be
+ − 460
explored via the lazy sequences mechanism). Given the code
565
+ − 461
\<close>
+ − 462
+ − 463
ML %grayML\<open>fun resolve_xmp_tac ctxt = resolve_tac ctxt [@{thm impI}, @{thm conjI}]\<close>
+ − 464
+ − 465
text \<open>
99
+ − 466
an example for @{ML resolve_tac} is the following proof where first an outermost
+ − 467
implication is analysed and then an outermost conjunction.
565
+ − 468
\<close>
99
+ − 469
362
+ − 470
lemma
+ − 471
shows "C \<longrightarrow> (A \<and> B)"
+ − 472
and "(A \<longrightarrow> B) \<and> C"
565
+ − 473
apply(tactic \<open>resolve_xmp_tac @{context} 1\<close>)
+ − 474
apply(tactic \<open>resolve_xmp_tac @{context} 2\<close>)
+ − 475
txt\<open>\begin{minipage}{\textwidth}
104
+ − 476
@{subgoals [display]}
565
+ − 477
\end{minipage}\<close>
99
+ − 478
(*<*)oops(*>*)
+ − 479
565
+ − 480
text \<open>
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 481
Similar versions taking a list of theorems exist for the tactics
562
+ − 482
@{ML_ind dresolve_tac in Tactic},
+ − 483
@{ML_ind eresolve_tac in Tactic} and so on.
363
+ − 484
+ − 485
Another simple tactic is @{ML_ind cut_facts_tac in Tactic}. It inserts a
+ − 486
list of theorems into the assumptions of the current goal state. Below we
+ − 487
will insert the definitions for the constants @{term True} and @{term
+ − 488
False}. So
565
+ − 489
\<close>
99
+ − 490
362
+ − 491
lemma
+ − 492
shows "True \<noteq> False"
565
+ − 493
apply(tactic \<open>cut_facts_tac [@{thm True_def}, @{thm False_def}] 1\<close>)
+ − 494
txt\<open>produces the goal state
368
+ − 495
\begin{isabelle}
107
+ − 496
@{subgoals [display]}
565
+ − 497
\end{isabelle}\<close>
107
+ − 498
(*<*)oops(*>*)
+ − 499
565
+ − 500
text \<open>
109
+ − 501
Often proofs on the ML-level involve elaborate operations on assumptions and
565
+ − 502
\<open>\<And>\<close>-quantified variables. To do such operations using the basic tactics
128
+ − 503
shown so far is very unwieldy and brittle. Some convenience and
363
+ − 504
safety is provided by the functions @{ML_ind FOCUS in Subgoal} and
+ − 505
@{ML_ind SUBPROOF in Subgoal}. These tactics fix the parameters
298
+ − 506
and bind the various components of a goal state to a record.
362
+ − 507
To see what happens, suppose the function defined in Figure~\ref{fig:sptac}, which
363
+ − 508
takes a record and just prints out the contents of this record. Then consider
+ − 509
the proof:
565
+ − 510
\<close>
+ − 511
+ − 512
+ − 513
text_raw\<open>
173
+ − 514
\begin{figure}[t]
177
+ − 515
\begin{minipage}{\textwidth}
99
+ − 516
\begin{isabelle}
565
+ − 517
\<close>
+ − 518
+ − 519
ML %grayML\<open>fun foc_tac {context, params, prems, asms, concl, schematics} =
132
+ − 520
let
440
+ − 521
fun assgn1 f1 f2 xs =
541
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 522
let
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 523
val out = map (fn (x, y) => Pretty.enum ":=" "" "" [f1 x, f2 y]) xs
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 524
in
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 525
Pretty.list "" "" out
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 526
end
440
+ − 527
562
+ − 528
fun assgn2 f xs = assgn1 (fn (n,T) => pretty_term context (Var (n,T))) f xs
440
+ − 529
+ − 530
val pps = map (fn (s, pp) => Pretty.block [Pretty.str s, pp])
+ − 531
[("params: ", assgn1 Pretty.str (pretty_cterm context) params),
+ − 532
("assumptions: ", pretty_cterms context asms),
+ − 533
("conclusion: ", pretty_cterm context concl),
+ − 534
("premises: ", pretty_thms_no_vars context prems),
+ − 535
("schematics: ", assgn2 (pretty_cterm context) (snd schematics))]
132
+ − 536
in
440
+ − 537
tracing (Pretty.string_of (Pretty.chunks pps)); all_tac
565
+ − 538
end\<close>
+ − 539
+ − 540
text_raw\<open>
99
+ − 541
\end{isabelle}
177
+ − 542
\end{minipage}
298
+ − 543
\caption{A function that prints out the various parameters provided by
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 544
@{ML FOCUS in Subgoal} and @{ML SUBPROOF}. It uses the functions defined
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 545
in Section~\ref{sec:printing} for extracting strings from @{ML_type cterm}s
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 546
and @{ML_type thm}s.\label{fig:sptac}}
99
+ − 547
\end{figure}
565
+ − 548
\<close>
95
+ − 549
562
+ − 550
schematic_goal
362
+ − 551
shows "\<And>x y. A x y \<Longrightarrow> B y x \<longrightarrow> C (?z y) x"
565
+ − 552
apply(tactic \<open>Subgoal.FOCUS foc_tac @{context} 1\<close>)
+ − 553
+ − 554
txt \<open>
109
+ − 555
The tactic produces the following printout:
95
+ − 556
99
+ − 557
\begin{quote}\small
95
+ − 558
\begin{tabular}{ll}
565
+ − 559
params: & \<open>x:= x\<close>, \<open>y:= y\<close>\\
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 560
assumptions: & @{term "A x y"}\\
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 561
conclusion: & @{term "B y x \<longrightarrow> C (z y) x"}\\
541
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 562
premises: & @{term "A x y"}\\
565
+ − 563
schematics: & \<open>?z:=z\<close>
95
+ − 564
\end{tabular}
99
+ − 565
\end{quote}
+ − 566
565
+ − 567
The \<open>params\<close> and \<open>schematics\<close> stand for list of pairs where the
+ − 568
left-hand side of \<open>:=\<close> is replaced by the right-hand side inside the
363
+ − 569
tactic. Notice that in the actual output the variables @{term x} and @{term
+ − 570
y} have a brown colour. Although they are parameters in the original goal,
+ − 571
they are fixed inside the tactic. By convention these fixed variables are
565
+ − 572
printed in brown colour. Similarly the schematic variable \<open>?z\<close>. The
363
+ − 573
assumption, or premise, @{prop "A x y"} is bound as @{ML_type cterm} to the
565
+ − 574
record-variable \<open>asms\<close>, but also as @{ML_type thm} to \<open>prems\<close>.
+ − 575
+ − 576
If we continue the proof script by applying the \<open>impI\<close>-rule
+ − 577
\<close>
95
+ − 578
+ − 579
apply(rule impI)
565
+ − 580
apply(tactic \<open>Subgoal.FOCUS foc_tac @{context} 1\<close>)
+ − 581
+ − 582
txt \<open>
118
+ − 583
then the tactic prints out:
95
+ − 584
99
+ − 585
\begin{quote}\small
95
+ − 586
\begin{tabular}{ll}
565
+ − 587
params: & \<open>x:= x\<close>, \<open>y:= y\<close>\\
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 588
assumptions: & @{term "A x y"}, @{term "B y x"}\\
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 589
conclusion: & @{term "C (z y) x"}\\
541
96d10631eec2
added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 590
premises: & @{term "A x y"}, @{term "B y x"}\\
565
+ − 591
schematics: & \<open>?z:=z\<close>
95
+ − 592
\end{tabular}
99
+ − 593
\end{quote}
565
+ − 594
\<close>
95
+ − 595
(*<*)oops(*>*)
+ − 596
565
+ − 597
text \<open>
+ − 598
Now also @{term "B y x"} is an assumption bound to \<open>asms\<close> and \<open>prems\<close>.
99
+ − 599
362
+ − 600
One difference between the tactics @{ML SUBPROOF} and @{ML FOCUS in Subgoal}
301
+ − 601
is that the former expects that the goal is solved completely, which the
363
+ − 602
latter does not. Another is that @{ML SUBPROOF} cannot instantiate any schematic
301
+ − 603
variables.
+ − 604
411
+ − 605
Observe that inside @{ML FOCUS in Subgoal} and @{ML SUBPROOF}, we are in a goal
562
+ − 606
state where there is only a conclusion. This means the tactics @{ML dresolve_tac} and
411
+ − 607
the like are of no use for manipulating the goal state. The assumptions inside
+ − 608
@{ML FOCUS in Subgoal} and @{ML SUBPROOF} are given as cterms and theorems in
565
+ − 609
the arguments \<open>asms\<close> and \<open>prems\<close>, respectively. This
411
+ − 610
means we can apply them using the usual assumption tactics. With this you can
562
+ − 611
for example easily implement a tactic that behaves almost like @{ML assume_tac}:
565
+ − 612
\<close>
+ − 613
+ − 614
ML %grayML\<open>val atac' = Subgoal.FOCUS (fn {prems, context, ...} => resolve_tac context prems 1)\<close>
+ − 615
+ − 616
text \<open>
109
+ − 617
If you apply @{ML atac'} to the next lemma
565
+ − 618
\<close>
107
+ − 619
362
+ − 620
lemma
+ − 621
shows "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y"
565
+ − 622
apply(tactic \<open>atac' @{context} 1\<close>)
+ − 623
txt\<open>it will produce
+ − 624
@{subgoals [display]}\<close>
99
+ − 625
(*<*)oops(*>*)
+ − 626
565
+ − 627
text \<open>
301
+ − 628
Notice that @{ML atac'} inside @{ML FOCUS in Subgoal} calls @{ML
565
+ − 629
resolve_tac} with the subgoal number \<open>1\<close> and also the outer call to
+ − 630
@{ML FOCUS in Subgoal} in the \isacommand{apply}-step uses \<open>1\<close>. This
301
+ − 631
is another advantage of @{ML FOCUS in Subgoal} and @{ML SUBPROOF}: the
+ − 632
addressing inside it is completely local to the tactic inside the
+ − 633
subproof. It is therefore possible to also apply @{ML atac'} to the second
+ − 634
goal by just writing:
104
+ − 635
565
+ − 636
\<close>
104
+ − 637
362
+ − 638
lemma
+ − 639
shows "True"
+ − 640
and "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y"
565
+ − 641
apply(tactic \<open>atac' @{context} 2\<close>)
105
+ − 642
apply(rule TrueI)
+ − 643
done
104
+ − 644
565
+ − 645
text \<open>
451
+ − 646
To sum up, both @{ML FOCUS in Subgoal} and @{ML SUBPROOF} are rather
411
+ − 647
convenient, but can impose a considerable run-time penalty in automatic
+ − 648
proofs. If speed is of the essence, then maybe the two lower level combinators
+ − 649
described next are more appropriate.
+ − 650
+ − 651
105
+ − 652
\begin{readmore}
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 653
The functions @{ML FOCUS in Subgoal} and @{ML SUBPROOF} are defined in
562
+ − 654
@{ML_file "Pure/Isar/subgoal.ML"} and also described in
298
+ − 655
\isccite{sec:results}.
105
+ − 656
\end{readmore}
+ − 657
411
+ − 658
Similar but less powerful functions than @{ML FOCUS in Subgoal},
+ − 659
respectively @{ML SUBPROOF}, are @{ML_ind SUBGOAL in Tactical} and @{ML_ind
+ − 660
CSUBGOAL in Tactical}. They allow you to inspect a given subgoal (the former
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 661
presents the subgoal as a @{ML_type term}, while the latter as a @{ML_type
363
+ − 662
cterm}). With them you can implement a tactic that applies a rule according
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 663
to the topmost logic connective in the subgoal (to illustrate this we only
411
+ − 664
analyse a few connectives). The code of the tactic is as follows.
565
+ − 665
\<close>
+ − 666
+ − 667
ML %linenosgray\<open>fun select_tac ctxt (t, i) =
99
+ − 668
case t of
562
+ − 669
@{term "Trueprop"} $ t' => select_tac ctxt (t', i)
+ − 670
| @{term "(\<Longrightarrow>)"} $ _ $ t' => select_tac ctxt (t', i)
+ − 671
| @{term "(\<and>)"} $ _ $ _ => resolve_tac ctxt [@{thm conjI}] i
+ − 672
| @{term "(\<longrightarrow>)"} $ _ $ _ => resolve_tac ctxt [@{thm impI}] i
+ − 673
| @{term "Not"} $ _ => resolve_tac ctxt [@{thm notI}] i
+ − 674
| Const (@{const_name "All"}, _) $ _ => resolve_tac ctxt [@{thm allI}] i
565
+ − 675
| _ => all_tac\<close>text_raw\<open>\label{tac:selecttac}\<close>
+ − 676
+ − 677
text \<open>
109
+ − 678
The input of the function is a term representing the subgoal and a number
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 679
specifying the subgoal of interest. In Line 3 you need to descend under the
109
+ − 680
outermost @{term "Trueprop"} in order to get to the connective you like to
+ − 681
analyse. Otherwise goals like @{prop "A \<and> B"} are not properly
+ − 682
analysed. Similarly with meta-implications in the next line. While for the
565
+ − 683
first five patterns we can use the \<open>@term\<close>-antiquotation to
109
+ − 684
construct the patterns, the pattern in Line 8 cannot be constructed in this
+ − 685
way. The reason is that an antiquotation would fix the type of the
363
+ − 686
quantified variable. So you really have to construct this pattern using the
+ − 687
basic term-constructors. This is not necessary in the other cases, because their
156
+ − 688
type is always fixed to function types involving only the type @{typ
298
+ − 689
bool}. (See Section \ref{sec:terms_types_manually} about constructing terms
156
+ − 690
manually.) For the catch-all pattern, we chose to just return @{ML all_tac}.
+ − 691
Consequently, @{ML select_tac} never fails.
+ − 692
105
+ − 693
107
+ − 694
Let us now see how to apply this tactic. Consider the four goals:
565
+ − 695
\<close>
105
+ − 696
+ − 697
362
+ − 698
lemma
+ − 699
shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F"
565
+ − 700
apply(tactic \<open>SUBGOAL (select_tac @{context}) 4\<close>)
+ − 701
apply(tactic \<open>SUBGOAL (select_tac @{context}) 3\<close>)
+ − 702
apply(tactic \<open>SUBGOAL (select_tac @{context}) 2\<close>)
+ − 703
apply(tactic \<open>SUBGOAL (select_tac @{context}) 1\<close>)
+ − 704
txt\<open>\begin{minipage}{\textwidth}
107
+ − 705
@{subgoals [display]}
565
+ − 706
\end{minipage}\<close>
99
+ − 707
(*<*)oops(*>*)
+ − 708
565
+ − 709
text \<open>
363
+ − 710
where in all but the last the tactic applies an introduction rule.
109
+ − 711
Note that we applied the tactic to the goals in ``reverse'' order.
+ − 712
This is a trick in order to be independent from the subgoals that are
+ − 713
produced by the rule. If we had applied it in the other order
565
+ − 714
\<close>
105
+ − 715
362
+ − 716
lemma
+ − 717
shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F"
565
+ − 718
apply(tactic \<open>SUBGOAL (select_tac @{context}) 1\<close>)
+ − 719
apply(tactic \<open>SUBGOAL (select_tac @{context}) 3\<close>)
+ − 720
apply(tactic \<open>SUBGOAL (select_tac @{context}) 4\<close>)
+ − 721
apply(tactic \<open>SUBGOAL (select_tac @{context}) 5\<close>)
105
+ − 722
(*<*)oops(*>*)
99
+ − 723
565
+ − 724
text \<open>
109
+ − 725
then we have to be careful to not apply the tactic to the two subgoals produced by
+ − 726
the first goal. To do this can result in quite messy code. In contrast,
107
+ − 727
the ``reverse application'' is easy to implement.
104
+ − 728
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 729
Of course, this example is
149
+ − 730
contrived: there are much simpler methods available in Isabelle for
243
+ − 731
implementing a tactic analysing a goal according to its topmost
149
+ − 732
connective. These simpler methods use tactic combinators, which we will
363
+ − 733
explain in the next section. But before that we will show how
+ − 734
tactic application can be constrained.
362
+ − 735
368
+ − 736
\begin{readmore}
+ − 737
The functions @{ML SUBGOAL} and @{ML CSUBGOAL} are defined in @{ML_file "Pure/tactical.ML"}.
+ − 738
\end{readmore}
+ − 739
+ − 740
562
+ − 741
Since @{ML_ind resolve_tac in Tactic} and the like use higher-order unification, an
368
+ − 742
automatic proof procedure based on them might become too fragile, if it just
+ − 743
applies theorems as shown above. The reason is that a number of theorems
+ − 744
introduce schematic variables into the goal state. Consider for example the
+ − 745
proof
565
+ − 746
\<close>
362
+ − 747
+ − 748
lemma
+ − 749
shows "\<forall>x \<in> A. P x \<Longrightarrow> Q x"
565
+ − 750
apply(tactic \<open>dresolve_tac @{context} [@{thm bspec}] 1\<close>)
+ − 751
txt\<open>\begin{minipage}{\textwidth}
362
+ − 752
@{subgoals [display]}
565
+ − 753
\end{minipage}\<close>
362
+ − 754
(*<*)oops(*>*)
+ − 755
565
+ − 756
text \<open>
+ − 757
where the application of theorem \<open>bspec\<close> generates two subgoals involving the
+ − 758
new schematic variable \<open>?x\<close>. Now, if you are not careful, tactics
363
+ − 759
applied to the first subgoal might instantiate this schematic variable in such a
565
+ − 760
way that the second subgoal becomes unprovable. If it is clear what the \<open>?x\<close>
362
+ − 761
should be, then this situation can be avoided by introducing a more
565
+ − 762
constrained version of the \<open>bspec\<close>-theorem. One way to give such
362
+ − 763
constraints is by pre-instantiating theorems with other theorems.
363
+ − 764
The function @{ML_ind RS in Drule}, for example, does this.
362
+ − 765
+ − 766
@{ML_response_fake [display,gray]
+ − 767
"@{thm disjI1} RS @{thm conjI}" "\<lbrakk>?P1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<or> ?Q1) \<and> ?Q"}
+ − 768
565
+ − 769
In this example it instantiates the first premise of the \<open>conjI\<close>-theorem
+ − 770
with the theorem \<open>disjI1\<close>. If the instantiation is impossible, as in the
362
+ − 771
case of
+ − 772
+ − 773
@{ML_response_fake_both [display,gray]
+ − 774
"@{thm conjI} RS @{thm mp}"
+ − 775
"*** Exception- THM (\"RSN: no unifiers\", 1,
+ − 776
[\"\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q\", \"\<lbrakk>?P \<longrightarrow> ?Q; ?P\<rbrakk> \<Longrightarrow> ?Q\"]) raised"}
+ − 777
363
+ − 778
then the function raises an exception. The function @{ML_ind RSN in Drule}
+ − 779
is similar to @{ML RS}, but takes an additional number as argument. This
+ − 780
number makes explicit which premise should be instantiated.
362
+ − 781
+ − 782
If you want to instantiate more than one premise of a theorem, you can use
363
+ − 783
the function @{ML_ind MRS in Drule}:
362
+ − 784
+ − 785
@{ML_response_fake [display,gray]
466
+ − 786
"[@{thm disjI1}, @{thm disjI2}] MRS @{thm conjI}"
+ − 787
"\<lbrakk>?P2; ?Q1\<rbrakk> \<Longrightarrow> (?P2 \<or> ?Q2) \<and> (?P1 \<or> ?Q1)"}
362
+ − 788
+ − 789
If you need to instantiate lists of theorems, you can use the
510
+ − 790
functions @{ML_ind RL in Drule} and @{ML_ind OF in Drule}. For
363
+ − 791
example in the code below, every theorem in the second list is
+ − 792
instantiated with every theorem in the first.
362
+ − 793
+ − 794
@{ML_response_fake [display,gray]
363
+ − 795
"let
+ − 796
val list1 = [@{thm impI}, @{thm disjI2}]
+ − 797
val list2 = [@{thm conjI}, @{thm disjI1}]
+ − 798
in
466
+ − 799
list1 RL list2
363
+ − 800
end"
466
+ − 801
"[\<lbrakk>?P1 \<Longrightarrow> ?Q1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<longrightarrow> ?Q1) \<and> ?Q,
+ − 802
\<lbrakk>?Q1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<or> ?Q1) \<and> ?Q,
+ − 803
(?P1 \<Longrightarrow> ?Q1) \<Longrightarrow> (?P1 \<longrightarrow> ?Q1) \<or> ?Q,
+ − 804
?Q1 \<Longrightarrow> (?P1 \<or> ?Q1) \<or> ?Q]"}
362
+ − 805
+ − 806
\begin{readmore}
368
+ − 807
The combinators for instantiating theorems with other theorems are
+ − 808
defined in @{ML_file "Pure/drule.ML"}.
362
+ − 809
\end{readmore}
+ − 810
+ − 811
Higher-order unification is also often in the way when applying certain
+ − 812
congruence theorems. For example we would expect that the theorem
+ − 813
@{thm [source] cong}
+ − 814
+ − 815
\begin{isabelle}
+ − 816
\isacommand{thm}~@{thm [source] cong}\\
565
+ − 817
\<open>> \<close>~@{thm cong}
362
+ − 818
\end{isabelle}
+ − 819
+ − 820
is applicable in the following proof producing the subgoals
+ − 821
@{term "t x = s u"} and @{term "y = w"}.
565
+ − 822
\<close>
362
+ − 823
+ − 824
lemma
+ − 825
fixes x y u w::"'a"
+ − 826
shows "t x y = s u w"
+ − 827
apply(rule cong)
565
+ − 828
txt\<open>\begin{minipage}{\textwidth}
362
+ − 829
@{subgoals [display]}
565
+ − 830
\end{minipage}\<close>
362
+ − 831
(*<*)oops(*>*)
+ − 832
565
+ − 833
text \<open>
363
+ − 834
As you can see this is unfortunately \emph{not} the case if we apply @{thm [source]
565
+ − 835
cong} with the method \<open>rule\<close>. The problem is
362
+ − 836
that higher-order unification produces an instantiation that is not the
+ − 837
intended one. While we can use \isacommand{back} to interactively find the
+ − 838
intended instantiation, this is not an option for an automatic proof
+ − 839
procedure. Fortunately, the tactic @{ML_ind cong_tac in Cong_Tac} helps
368
+ − 840
with applying congruence theorems and finding the intended instantiation.
363
+ − 841
For example
565
+ − 842
\<close>
362
+ − 843
+ − 844
lemma
+ − 845
fixes x y u w::"'a"
+ − 846
shows "t x y = s u w"
565
+ − 847
apply(tactic \<open>Cong_Tac.cong_tac @{context} @{thm cong} 1\<close>)
+ − 848
txt\<open>\begin{minipage}{\textwidth}
362
+ − 849
@{subgoals [display]}
565
+ − 850
\end{minipage}\<close>
362
+ − 851
(*<*)oops(*>*)
+ − 852
565
+ − 853
text \<open>
368
+ − 854
However, frequently it is necessary to explicitly match a theorem against a goal
363
+ − 855
state and in doing so construct manually an appropriate instantiation. Imagine
362
+ − 856
you have the theorem
565
+ − 857
\<close>
362
+ − 858
+ − 859
lemma rule:
+ − 860
shows "\<lbrakk>f = g; x = y\<rbrakk> \<Longrightarrow> R (f x) (g y)"
+ − 861
sorry
+ − 862
565
+ − 863
text \<open>
551
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 864
and you want to apply it to the goal @{term "(t\<^sub>1 t\<^sub>2) \<le>
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 865
(s\<^sub>1 s\<^sub>2)"}. Since in the theorem all variables are
565
+ − 866
schematic, we have a nasty higher-order unification problem and \<open>rtac\<close>
363
+ − 867
will not be able to apply this rule in the way we want. For the goal at hand
565
+ − 868
we want to use it so that @{term R} is instantiated to the constant \<open>\<le>\<close> and
363
+ − 869
similarly ``obvious'' instantiations for the other variables. To achieve this
+ − 870
we need to match the conclusion of
+ − 871
@{thm [source] rule} against the goal reading off an instantiation for
362
+ − 872
@{thm [source] rule}. For this the function @{ML_ind first_order_match in Thm}
363
+ − 873
matches two @{ML_type cterm}s and produces, in the successful case, a matcher
+ − 874
that can be used to instantiate the theorem. The instantiation
465
+ − 875
can be done with the function @{ML_ind instantiate_normalize in Drule}. To automate
362
+ − 876
this we implement the following function.
565
+ − 877
\<close>
+ − 878
+ − 879
ML %linenosgray\<open>fun fo_rtac thm = Subgoal.FOCUS (fn {context, concl, ...} =>
362
+ − 880
let
562
+ − 881
val concl_pat = Drule.strip_imp_concl (Thm.cprop_of thm)
362
+ − 882
val insts = Thm.first_order_match (concl_pat, concl)
+ − 883
in
562
+ − 884
resolve_tac context [(Drule.instantiate_normalize insts thm)] 1
565
+ − 885
end)\<close>
+ − 886
+ − 887
text \<open>
363
+ − 888
Note that we use @{ML FOCUS in Subgoal} because it gives us directly access
+ − 889
to the conclusion of the goal state, but also because this function
+ − 890
takes care of correctly handling parameters that might be present
+ − 891
in the goal state. In Line 3 we use the function @{ML_ind strip_imp_concl in Drule}
+ − 892
for calculating the conclusion of a theorem (produced as @{ML_type cterm}).
+ − 893
An example of @{ML fo_rtac} is as follows.
565
+ − 894
\<close>
362
+ − 895
+ − 896
lemma
551
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 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)"
565
+ − 898
apply(tactic \<open>fo_rtac @{thm rule} @{context} 1\<close>)
+ − 899
txt\<open>\begin{minipage}{\textwidth}
362
+ − 900
@{subgoals [display]}
565
+ − 901
\end{minipage}\<close>
362
+ − 902
(*<*)oops(*>*)
+ − 903
565
+ − 904
text \<open>
362
+ − 905
We obtain the intended subgoals and also the parameters are correctly
+ − 906
introduced in both of them. Such manual instantiations are quite frequently
368
+ − 907
necessary in order to appropriately constrain the application of theorems.
+ − 908
Otherwise one can end up with ``wild'' higher-order unification problems
+ − 909
that make automatic proofs fail.
+ − 910
+ − 911
\begin{readmore}
+ − 912
Functions for matching @{ML_type cterm}s are defined in @{ML_file "Pure/thm.ML"}.
+ − 913
Functions for instantiating schematic variables in theorems are defined
+ − 914
in @{ML_file "Pure/drule.ML"}.
+ − 915
\end{readmore}
565
+ − 916
\<close>
+ − 917
+ − 918
section \<open>Tactic Combinators\<close>
+ − 919
+ − 920
text \<open>
109
+ − 921
The purpose of tactic combinators is to build compound tactics out of
363
+ − 922
smaller tactics. In the previous section we already used @{ML_ind THEN in Tactical},
+ − 923
which just strings together two tactics in a sequence. For example:
565
+ − 924
\<close>
93
+ − 925
362
+ − 926
lemma
+ − 927
shows "(Foo \<and> Bar) \<and> False"
565
+ − 928
apply(tactic \<open>resolve_tac @{context} [@{thm conjI}] 1 THEN resolve_tac @{context} [@{thm conjI}] 1\<close>)
+ − 929
txt \<open>\begin{minipage}{\textwidth}
105
+ − 930
@{subgoals [display]}
565
+ − 931
\end{minipage}\<close>
105
+ − 932
(*<*)oops(*>*)
99
+ − 933
565
+ − 934
text \<open>
368
+ − 935
If you want to avoid the hard-coded subgoal addressing in each component,
+ − 936
then, as seen earlier, you can use the ``primed'' version of @{ML THEN}.
+ − 937
For example:
565
+ − 938
\<close>
93
+ − 939
362
+ − 940
lemma
+ − 941
shows "(Foo \<and> Bar) \<and> False"
565
+ − 942
apply(tactic \<open>(resolve_tac @{context} [@{thm conjI}] THEN' resolve_tac @{context} [@{thm conjI}]) 1\<close>)
+ − 943
txt \<open>\begin{minipage}{\textwidth}
105
+ − 944
@{subgoals [display]}
565
+ − 945
\end{minipage}\<close>
93
+ − 946
(*<*)oops(*>*)
+ − 947
565
+ − 948
text \<open>
404
+ − 949
Here you have to specify the subgoal of interest only once and it is
+ − 950
consistently applied to the component. For most tactic combinators such a
+ − 951
``primed'' version exists and in what follows we will usually prefer it over
+ − 952
the ``unprimed'' one.
+ − 953
565
+ − 954
The tactic combinator @{ML_ind RANGE in Tactical} takes a list of \<open>n\<close>
+ − 955
tactics, say, and applies them to each of the first \<open>n\<close> subgoals.
404
+ − 956
For example below we first apply the introduction rule for conjunctions
+ − 957
and then apply a tactic to each of the two subgoals.
565
+ − 958
\<close>
404
+ − 959
+ − 960
lemma
+ − 961
shows "A \<Longrightarrow> True \<and> A"
565
+ − 962
apply(tactic \<open>(resolve_tac @{context} [@{thm conjI}]
+ − 963
THEN' RANGE [resolve_tac @{context} [@{thm TrueI}], assume_tac @{context}]) 1\<close>)
+ − 964
txt \<open>\begin{minipage}{\textwidth}
404
+ − 965
@{subgoals [display]}
565
+ − 966
\end{minipage}\<close>
404
+ − 967
(*<*)oops(*>*)
+ − 968
565
+ − 969
text \<open>
404
+ − 970
If there is a list of tactics that should all be tried out in sequence on
+ − 971
one specified subgoal, you can use the combinator @{ML_ind EVERY' in
+ − 972
Tactical}. For example the function @{ML foo_tac'} from
+ − 973
page~\pageref{tac:footacprime} can also be written as:
565
+ − 974
\<close>
+ − 975
+ − 976
ML %grayML\<open>fun foo_tac'' ctxt =
562
+ − 977
EVERY' [eresolve_tac ctxt [@{thm disjE}], resolve_tac ctxt [@{thm disjI2}],
565
+ − 978
assume_tac ctxt, resolve_tac ctxt [@{thm disjI1}], assume_tac ctxt]\<close>
+ − 979
+ − 980
text \<open>
109
+ − 981
There is even another way of implementing this tactic: in automatic proof
+ − 982
procedures (in contrast to tactics that might be called by the user) there
+ − 983
are often long lists of tactics that are applied to the first
562
+ − 984
subgoal. Instead of writing the code above and then calling @{ML "foo_tac'' @{context} 1"},
109
+ − 985
you can also just write
565
+ − 986
\<close>
+ − 987
+ − 988
ML %grayML\<open>fun foo_tac1 ctxt =
562
+ − 989
EVERY1 [eresolve_tac ctxt [@{thm disjE}], resolve_tac ctxt [@{thm disjI2}],
565
+ − 990
assume_tac ctxt, resolve_tac ctxt [@{thm disjI1}], assume_tac ctxt]\<close>
+ − 991
+ − 992
text \<open>
118
+ − 993
and call @{ML foo_tac1}.
109
+ − 994
363
+ − 995
With the combinators @{ML THEN'}, @{ML EVERY'} and @{ML_ind EVERY1 in Tactical} it must be
109
+ − 996
guaranteed that all component tactics successfully apply; otherwise the
+ − 997
whole tactic will fail. If you rather want to try out a number of tactics,
363
+ − 998
then you can use the combinator @{ML_ind ORELSE' in Tactical} for two tactics, and @{ML_ind
+ − 999
FIRST' in Tactical} (or @{ML_ind FIRST1 in Tactical}) for a list of tactics. For example, the tactic
109
+ − 1000
565
+ − 1001
\<close>
+ − 1002
+ − 1003
ML %grayML\<open>fun orelse_xmp_tac ctxt =
+ − 1004
resolve_tac ctxt [@{thm disjI1}] ORELSE' resolve_tac ctxt [@{thm conjI}]\<close>
+ − 1005
+ − 1006
text \<open>
+ − 1007
will first try out whether theorem \<open>disjI\<close> applies and in case of failure
+ − 1008
will try \<open>conjI\<close>. To see this consider the proof
+ − 1009
\<close>
105
+ − 1010
362
+ − 1011
lemma
+ − 1012
shows "True \<and> False" and "Foo \<or> Bar"
565
+ − 1013
apply(tactic \<open>orelse_xmp_tac @{context} 2\<close>)
+ − 1014
apply(tactic \<open>orelse_xmp_tac @{context} 1\<close>)
+ − 1015
txt \<open>which results in the goal state
368
+ − 1016
\begin{isabelle}
107
+ − 1017
@{subgoals [display]}
368
+ − 1018
\end{isabelle}
565
+ − 1019
\<close>
93
+ − 1020
(*<*)oops(*>*)
+ − 1021
+ − 1022
565
+ − 1023
text \<open>
109
+ − 1024
Using @{ML FIRST'} we can simplify our @{ML select_tac} from Page~\pageref{tac:selecttac}
+ − 1025
as follows:
565
+ − 1026
\<close>
+ − 1027
+ − 1028
ML %grayML\<open>fun select_tac' ctxt =
562
+ − 1029
FIRST' [resolve_tac ctxt [@{thm conjI}], resolve_tac ctxt [@{thm impI}],
565
+ − 1030
resolve_tac ctxt [@{thm notI}], resolve_tac ctxt [@{thm allI}], K all_tac]\<close>text_raw\<open>\label{tac:selectprime}\<close>
+ − 1031
+ − 1032
text \<open>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1033
Since we like to mimic the behaviour of @{ML select_tac} as closely as possible,
109
+ − 1034
we must include @{ML all_tac} at the end of the list, otherwise the tactic will
368
+ − 1035
fail if no theorem applies (we also have to wrap @{ML all_tac} using the
109
+ − 1036
@{ML K}-combinator, because it does not take a subgoal number as argument). You can
+ − 1037
test the tactic on the same goals:
565
+ − 1038
\<close>
107
+ − 1039
362
+ − 1040
lemma
+ − 1041
shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F"
565
+ − 1042
apply(tactic \<open>select_tac' @{context} 4\<close>)
+ − 1043
apply(tactic \<open>select_tac' @{context} 3\<close>)
+ − 1044
apply(tactic \<open>select_tac' @{context} 2\<close>)
+ − 1045
apply(tactic \<open>select_tac' @{context} 1\<close>)
+ − 1046
txt\<open>\begin{minipage}{\textwidth}
107
+ − 1047
@{subgoals [display]}
565
+ − 1048
\end{minipage}\<close>
107
+ − 1049
(*<*)oops(*>*)
+ − 1050
565
+ − 1051
text \<open>
109
+ − 1052
Since such repeated applications of a tactic to the reverse order of
+ − 1053
\emph{all} subgoals is quite common, there is the tactic combinator
363
+ − 1054
@{ML_ind ALLGOALS in Tactical} that simplifies this. Using this combinator you can simply
565
+ − 1055
write:\<close>
107
+ − 1056
362
+ − 1057
lemma
+ − 1058
shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F"
565
+ − 1059
apply(tactic \<open>ALLGOALS (select_tac' @{context})\<close>)
+ − 1060
txt\<open>\begin{minipage}{\textwidth}
107
+ − 1061
@{subgoals [display]}
565
+ − 1062
\end{minipage}\<close>
107
+ − 1063
(*<*)oops(*>*)
+ − 1064
565
+ − 1065
text \<open>
109
+ − 1066
Remember that we chose to implement @{ML select_tac'} so that it
368
+ − 1067
always succeeds by fact of having @{ML all_tac} at the end of the tactic
363
+ − 1068
list. The same can be achieved with the tactic combinator @{ML_ind TRY in Tactical}.
243
+ − 1069
For example:
565
+ − 1070
\<close>
+ − 1071
+ − 1072
ML %grayML\<open>fun select_tac'' ctxt =
562
+ − 1073
TRY o FIRST' [resolve_tac ctxt [@{thm conjI}], resolve_tac ctxt [@{thm impI}],
565
+ − 1074
resolve_tac ctxt [@{thm notI}], resolve_tac ctxt [@{thm allI}]]\<close>
+ − 1075
text_raw\<open>\label{tac:selectprimeprime}\<close>
+ − 1076
+ − 1077
text \<open>
243
+ − 1078
This tactic behaves in the same way as @{ML select_tac'}: it tries out
+ − 1079
one of the given tactics and if none applies leaves the goal state
+ − 1080
unchanged. This, however, can be potentially very confusing when visible to
+ − 1081
the user, for example, in cases where the goal is the form
+ − 1082
565
+ − 1083
\<close>
107
+ − 1084
362
+ − 1085
lemma
+ − 1086
shows "E \<Longrightarrow> F"
565
+ − 1087
apply(tactic \<open>select_tac' @{context} 1\<close>)
+ − 1088
txt\<open>\begin{minipage}{\textwidth}
107
+ − 1089
@{subgoals [display]}
565
+ − 1090
\end{minipage}\<close>
107
+ − 1091
(*<*)oops(*>*)
+ − 1092
565
+ − 1093
text \<open>
368
+ − 1094
In this case no theorem applies. But because we wrapped the tactic in a @{ML
+ − 1095
TRY}, it does not fail anymore. The problem is that for the user there is
+ − 1096
little chance to see whether progress in the proof has been made, or not. By
+ − 1097
convention therefore, tactics visible to the user should either change
+ − 1098
something or fail.
+ − 1099
109
+ − 1100
To comply with this convention, we could simply delete the @{ML "K all_tac"}
368
+ − 1101
in @{ML select_tac'} or delete @{ML TRY} from @{ML select_tac''}. But for
+ − 1102
the sake of argument, let us suppose that this deletion is \emph{not} an
+ − 1103
option. In such cases, you can use the combinator @{ML_ind CHANGED in
+ − 1104
Tactical} to make sure the subgoal has been changed by the tactic. Because
+ − 1105
now
565
+ − 1106
\<close>
107
+ − 1107
362
+ − 1108
lemma
+ − 1109
shows "E \<Longrightarrow> F"
565
+ − 1110
apply(tactic \<open>CHANGED (select_tac' @{context} 1)\<close>)(*<*)?(*>*)
+ − 1111
txt\<open>gives the error message:
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1112
\begin{isabelle}
565
+ − 1113
\<open>*** empty result sequence -- proof command failed\<close>\\
+ − 1114
\<open>*** At command "apply".\<close>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1115
\end{isabelle}
565
+ − 1116
\<close>(*<*)oops(*>*)
+ − 1117
+ − 1118
+ − 1119
text \<open>
368
+ − 1120
We can further extend the @{ML select_tac}s so that they not just apply to the topmost
109
+ − 1121
connective, but also to the ones immediately ``underneath'', i.e.~analyse the goal
368
+ − 1122
completely. For this you can use the tactic combinator @{ML_ind REPEAT in Tactical}. As an example
109
+ − 1123
suppose the following tactic
565
+ − 1124
\<close>
+ − 1125
+ − 1126
ML %grayML\<open>fun repeat_xmp_tac ctxt = REPEAT (CHANGED (select_tac' ctxt 1))\<close>
+ − 1127
+ − 1128
text \<open>which applied to the proof\<close>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1129
362
+ − 1130
lemma
+ − 1131
shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)"
565
+ − 1132
apply(tactic \<open>repeat_xmp_tac @{context}\<close>)
+ − 1133
txt\<open>produces
368
+ − 1134
\begin{isabelle}
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1135
@{subgoals [display]}
565
+ − 1136
\end{isabelle}\<close>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1137
(*<*)oops(*>*)
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1138
565
+ − 1139
text \<open>
368
+ − 1140
Here it is crucial that @{ML select_tac'} is prefixed with @{ML CHANGED},
109
+ − 1141
because otherwise @{ML REPEAT} runs into an infinite loop (it applies the
368
+ − 1142
tactic as long as it succeeds). The function @{ML_ind REPEAT1 in Tactical}
+ − 1143
is similar, but runs the tactic at least once (failing if this is not
+ − 1144
possible).
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1145
238
+ − 1146
If you are after the ``primed'' version of @{ML repeat_xmp_tac}, then you
243
+ − 1147
can implement it as
565
+ − 1148
\<close>
+ − 1149
+ − 1150
ML %grayML\<open>fun repeat_xmp_tac' ctxt = REPEAT o CHANGED o select_tac' ctxt\<close>
+ − 1151
+ − 1152
text \<open>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1153
since there are no ``primed'' versions of @{ML REPEAT} and @{ML CHANGED}.
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1154
243
+ − 1155
If you look closely at the goal state above, then you see the tactics @{ML repeat_xmp_tac}
238
+ − 1156
and @{ML repeat_xmp_tac'} are not yet quite what we are after: the problem is
109
+ − 1157
that goals 2 and 3 are not analysed. This is because the tactic
+ − 1158
is applied repeatedly only to the first subgoal. To analyse also all
363
+ − 1159
resulting subgoals, you can use the tactic combinator @{ML_ind REPEAT_ALL_NEW in Tactical}.
368
+ − 1160
Supposing the tactic
565
+ − 1161
\<close>
+ − 1162
+ − 1163
ML %grayML\<open>fun repeat_all_new_xmp_tac ctxt = REPEAT_ALL_NEW (CHANGED o select_tac' ctxt)\<close>
+ − 1164
+ − 1165
text \<open>
368
+ − 1166
you can see that the following goal
565
+ − 1167
\<close>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1168
362
+ − 1169
lemma
+ − 1170
shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)"
565
+ − 1171
apply(tactic \<open>repeat_all_new_xmp_tac @{context} 1\<close>)
+ − 1172
txt\<open>\begin{minipage}{\textwidth}
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1173
@{subgoals [display]}
565
+ − 1174
\end{minipage}\<close>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1175
(*<*)oops(*>*)
93
+ − 1176
565
+ − 1177
text \<open>
109
+ − 1178
is completely analysed according to the theorems we chose to
120
+ − 1179
include in @{ML select_tac'}.
109
+ − 1180
+ − 1181
Recall that tactics produce a lazy sequence of successor goal states. These
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1182
states can be explored using the command \isacommand{back}. For example
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1183
565
+ − 1184
\<close>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1185
362
+ − 1186
lemma
+ − 1187
shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R"
565
+ − 1188
apply(tactic \<open>eresolve_tac @{context} [@{thm disjE}] 1\<close>)
+ − 1189
txt\<open>applies the rule to the first assumption yielding the goal state:
368
+ − 1190
\begin{isabelle}
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1191
@{subgoals [display]}
368
+ − 1192
\end{isabelle}
109
+ − 1193
+ − 1194
After typing
565
+ − 1195
\<close>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1196
(*<*)
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1197
oops
362
+ − 1198
lemma
+ − 1199
shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R"
565
+ − 1200
apply(tactic \<open>eresolve_tac @{context} [@{thm disjE}] 1\<close>)
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1201
(*>*)
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1202
back
565
+ − 1203
txt\<open>the rule now applies to the second assumption.
368
+ − 1204
\begin{isabelle}
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1205
@{subgoals [display]}
565
+ − 1206
\end{isabelle}\<close>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1207
(*<*)oops(*>*)
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1208
565
+ − 1209
text \<open>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1210
Sometimes this leads to confusing behaviour of tactics and also has
109
+ − 1211
the potential to explode the search space for tactics.
+ − 1212
These problems can be avoided by prefixing the tactic with the tactic
363
+ − 1213
combinator @{ML_ind DETERM in Tactical}.
565
+ − 1214
\<close>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1215
362
+ − 1216
lemma
+ − 1217
shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R"
565
+ − 1218
apply(tactic \<open>DETERM (eresolve_tac @{context} [@{thm disjE}] 1)\<close>)
+ − 1219
txt \<open>\begin{minipage}{\textwidth}
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1220
@{subgoals [display]}
565
+ − 1221
\end{minipage}\<close>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1222
(*<*)oops(*>*)
565
+ − 1223
text \<open>
118
+ − 1224
This combinator will prune the search space to just the first successful application.
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1225
Attempting to apply \isacommand{back} in this goal states gives the
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1226
error message:
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1227
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1228
\begin{isabelle}
565
+ − 1229
\<open>*** back: no alternatives\<close>\\
+ − 1230
\<open>*** At command "back".\<close>
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1231
\end{isabelle}
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1232
368
+ − 1233
369
+ − 1234
\footnote{\bf FIXME: say something about @{ML_ind COND in Tactical} and COND'}
368
+ − 1235
\footnote{\bf FIXME: PARALLEL-CHOICE PARALLEL-GOALS}
238
+ − 1236
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1237
\begin{readmore}
289
+ − 1238
Most tactic combinators described in this section are defined in @{ML_file "Pure/tactical.ML"}.
238
+ − 1239
Some combinators for the purpose of proof search are implemented in @{ML_file "Pure/search.ML"}.
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1240
\end{readmore}
565
+ − 1241
\<close>
+ − 1242
+ − 1243
text \<open>
313
+ − 1244
\begin{exercise}\label{ex:dyckhoff}
370
+ − 1245
Dyckhoff presents in \cite{Dyckhoff92} inference rules of a sequent
368
+ − 1246
calculus, called G4ip, for intuitionistic propositional logic. The
+ − 1247
interesting feature of this calculus is that no contraction rule is needed
370
+ − 1248
in order to be complete. As a result the rules can be applied exhaustively, which
+ − 1249
in turn leads to simple decision procedure for propositional intuitionistic logic.
+ − 1250
The task is to implement this decision procedure as a tactic. His rules are
314
+ − 1251
+ − 1252
\begin{center}
+ − 1253
\def\arraystretch{2.3}
+ − 1254
\begin{tabular}{cc}
+ − 1255
\infer[Ax]{A,\varGamma \Rightarrow A}{} &
+ − 1256
\infer[False]{False,\varGamma \Rightarrow G}{}\\
+ − 1257
+ − 1258
\infer[\wedge_L]{A \wedge B, \varGamma \Rightarrow G}{A, B, \varGamma \Rightarrow G} &
+ − 1259
\infer[\wedge_R]
+ − 1260
{\varGamma \Rightarrow A\wedge B}{\varGamma \Rightarrow A & \varGamma \Rightarrow B}\\
313
+ − 1261
314
+ − 1262
\infer[\vee_L]
+ − 1263
{A\vee B, \varGamma \Rightarrow G}{A,\varGamma \Rightarrow G & B,\varGamma \Rightarrow G} &
+ − 1264
\infer[\vee_{R_1}]
+ − 1265
{\varGamma \Rightarrow A \vee B}{\varGamma \Rightarrow A} \hspace{3mm}
+ − 1266
\infer[\vee_{R_2}]
+ − 1267
{\varGamma \Rightarrow A \vee B}{\varGamma \Rightarrow B}\\
+ − 1268
+ − 1269
\infer[\longrightarrow_{L_1}]
+ − 1270
{A\longrightarrow B, A, \varGamma \Rightarrow G}{B, A, \varGamma \Rightarrow G} &
+ − 1271
\infer[\longrightarrow_R]
+ − 1272
{\varGamma \Rightarrow A\longrightarrow B}{A,\varGamma \Rightarrow B}\\
+ − 1273
+ − 1274
\infer[\longrightarrow_{L_2}]
+ − 1275
{(C \wedge D)\longrightarrow B, \varGamma \Rightarrow G}
+ − 1276
{C\longrightarrow (D \longrightarrow B), \varGamma \Rightarrow G} &
+ − 1277
+ − 1278
\infer[\longrightarrow_{L_3}]
+ − 1279
{(C \vee D)\longrightarrow B, \varGamma \Rightarrow G}
+ − 1280
{C\longrightarrow B, D\longrightarrow B, \varGamma \Rightarrow G}\\
+ − 1281
+ − 1282
\multicolumn{2}{c}{
+ − 1283
\infer[\longrightarrow_{L_4}]
+ − 1284
{(C \longrightarrow D)\longrightarrow B, \varGamma \Rightarrow G}
+ − 1285
{D\longrightarrow B, \varGamma \Rightarrow C \longrightarrow D &
+ − 1286
B, \varGamma \Rightarrow G}}\\
+ − 1287
\end{tabular}
+ − 1288
\end{center}
+ − 1289
370
+ − 1290
Note that in Isabelle right rules need to be implemented as
+ − 1291
introduction rule, the left rules as elimination rules. You have to to
+ − 1292
prove separate theorems corresponding to $\longrightarrow_{L_{1..4}}$. The
+ − 1293
tactic should explore all possibilites of applying these rules to a
+ − 1294
propositional formula until a goal state is reached in which all subgoals
+ − 1295
are discharged. For this you can use the tactic combinator @{ML_ind
+ − 1296
DEPTH_SOLVE in Search} in the structure @{ML_struct Search}.
368
+ − 1297
\end{exercise}
+ − 1298
+ − 1299
\begin{exercise}
370
+ − 1300
Add to the sequent calculus from the previous exercise also rules for
+ − 1301
equality and run your tactic on the de Bruijn formulae discussed
+ − 1302
in Exercise~\ref{ex:debruijn}.
313
+ − 1303
\end{exercise}
565
+ − 1304
\<close>
+ − 1305
+ − 1306
section \<open>Simplifier Tactics\label{sec:simplifier}\<close>
+ − 1307
+ − 1308
text \<open>
368
+ − 1309
A lot of convenience in reasoning with Isabelle derives from its
370
+ − 1310
powerful simplifier. We will describe it in this section. However,
+ − 1311
due to its complexity, we can mostly only scratch the surface.
368
+ − 1312
+ − 1313
The power of the simplifier is a strength and a weakness at the same time,
+ − 1314
because you can easily make the simplifier run into a loop and in general
+ − 1315
its behaviour can be difficult to predict. There is also a multitude of
+ − 1316
options that you can configure to change the behaviour of the
+ − 1317
simplifier. There are the following five main tactics behind the simplifier
+ − 1318
(in parentheses is their user-level counterpart):
152
+ − 1319
+ − 1320
\begin{isabelle}
157
+ − 1321
\begin{center}
152
+ − 1322
\begin{tabular}{l@ {\hspace{2cm}}l}
565
+ − 1323
@{ML_ind simp_tac in Simplifier} & \<open>(simp (no_asm))\<close> \\
+ − 1324
@{ML_ind asm_simp_tac in Simplifier} & \<open>(simp (no_asm_simp))\<close> \\
+ − 1325
@{ML_ind full_simp_tac in Simplifier} & \<open>(simp (no_asm_use))\<close> \\
+ − 1326
@{ML_ind asm_lr_simp_tac in Simplifier} & \<open>(simp (asm_lr))\<close> \\
+ − 1327
@{ML_ind asm_full_simp_tac in Simplifier} & \<open>(simp)\<close>
152
+ − 1328
\end{tabular}
157
+ − 1329
\end{center}
152
+ − 1330
\end{isabelle}
+ − 1331
370
+ − 1332
All these tactics take a simpset and an integer as argument (the latter as usual
162
+ − 1333
to specify the goal to be analysed). So the proof
565
+ − 1334
\<close>
152
+ − 1335
362
+ − 1336
lemma
+ − 1337
shows "Suc (1 + 2) < 3 + 2"
152
+ − 1338
apply(simp)
+ − 1339
done
+ − 1340
565
+ − 1341
text \<open>
157
+ − 1342
corresponds on the ML-level to the tactic
565
+ − 1343
\<close>
152
+ − 1344
362
+ − 1345
lemma
+ − 1346
shows "Suc (1 + 2) < 3 + 2"
565
+ − 1347
apply(tactic \<open>asm_full_simp_tac @{context} 1\<close>)
152
+ − 1348
done
+ − 1349
565
+ − 1350
text \<open>
162
+ − 1351
If the simplifier cannot make any progress, then it leaves the goal unchanged,
209
+ − 1352
i.e., does not raise any error message. That means if you use it to unfold a
162
+ − 1353
definition for a constant and this constant is not present in the goal state,
+ − 1354
you can still safely apply the simplifier.
152
+ − 1355
368
+ − 1356
\footnote{\bf FIXME: show rewriting of cterms}
308
+ − 1357
412
+ − 1358
There is one restriction you have to keep in mind when using the simplifier:
413
+ − 1359
it can only deal with rewriting rules whose left-hand sides are higher order
+ − 1360
pattern (see Section \ref{sec:univ} on unification). Whether a term is a pattern
+ − 1361
or not can be tested with the function @{ML_ind pattern in Pattern} from
+ − 1362
the structure @{ML_struct Pattern}. If a rule is not a pattern and you want
+ − 1363
to use it for rewriting, then you have to use simprocs or conversions, both
+ − 1364
of which we shall describe in the next section.
412
+ − 1365
162
+ − 1366
When using the simplifier, the crucial information you have to provide is
368
+ − 1367
the simpset. If this information is not handled with care, then, as
+ − 1368
mentioned above, the simplifier can easily run into a loop. Therefore a good
+ − 1369
rule of thumb is to use simpsets that are as minimal as possible. It might
+ − 1370
be surprising that a simpset is more complex than just a simple collection
+ − 1371
of theorems. One reason for the complexity is that the simplifier must be
+ − 1372
able to rewrite inside terms and should also be able to rewrite according to
+ − 1373
theorems that have premises.
+ − 1374
+ − 1375
The rewriting inside terms requires congruence theorems, which
+ − 1376
are typically meta-equalities of the form
152
+ − 1377
+ − 1378
\begin{isabelle}
157
+ − 1379
\begin{center}
565
+ − 1380
\mbox{\inferrule{\<open>t\<^sub>1 \<equiv> s\<^sub>1 \<dots> t\<^sub>n \<equiv> s\<^sub>n\<close>}
+ − 1381
{\<open>constr t\<^sub>1\<dots>t\<^sub>n \<equiv> constr s\<^sub>1\<dots>s\<^sub>n\<close>}}
157
+ − 1382
\end{center}
152
+ − 1383
\end{isabelle}
+ − 1384
565
+ − 1385
with \<open>constr\<close> being a constant, like @{const "If"}, @{const "Let"}
368
+ − 1386
and so on. Every simpset contains only one congruence rule for each
+ − 1387
term-constructor, which however can be overwritten. The user can declare
565
+ − 1388
lemmas to be congruence rules using the attribute \<open>[cong]\<close>. Note that
370
+ − 1389
in HOL these congruence theorems are usually stated as equations, which are
+ − 1390
then internally transformed into meta-equations.
368
+ − 1391
+ − 1392
The rewriting with theorems involving premises requires what is in Isabelle
+ − 1393
called a subgoaler, a solver and a looper. These can be arbitrary tactics
+ − 1394
that can be installed in a simpset and which are executed at various stages
370
+ − 1395
during simplification.
+ − 1396
+ − 1397
Simpsets can also include simprocs, which produce rewrite rules on
+ − 1398
demand according to a pattern (see next section for a detailed description
+ − 1399
of simpsets). Another component are split-rules, which can simplify for
+ − 1400
example the ``then'' and ``else'' branches of if-statements under the
+ − 1401
corresponding preconditions.
162
+ − 1402
157
+ − 1403
\begin{readmore}
458
+ − 1404
For more information about the simplifier see @{ML_file "Pure/raw_simplifier.ML"}
368
+ − 1405
and @{ML_file "Pure/simplifier.ML"}. The generic splitter is implemented in
+ − 1406
@{ML_file "Provers/splitter.ML"}.
157
+ − 1407
\end{readmore}
152
+ − 1408
368
+ − 1409
+ − 1410
\footnote{\bf FIXME: Find the right place to mention this: Discrimination
+ − 1411
nets are implemented in @{ML_file "Pure/net.ML"}.}
160
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1412
370
+ − 1413
The most common combinators for modifying simpsets are:
152
+ − 1414
+ − 1415
\begin{isabelle}
370
+ − 1416
\begin{tabular}{l@ {\hspace{10mm}}l}
458
+ − 1417
@{ML_ind addsimps in Raw_Simplifier} & @{ML_ind delsimps in Raw_Simplifier}\\
+ − 1418
@{ML_ind addsimprocs in Raw_Simplifier} & @{ML_ind delsimprocs in Raw_Simplifier}\\
503
+ − 1419
@{ML_ind add_cong in Raw_Simplifier} & @{ML_ind del_cong in Raw_Simplifier}\\
152
+ − 1420
\end{tabular}
+ − 1421
\end{isabelle}
+ − 1422
565
+ − 1423
\<close>
+ − 1424
+ − 1425
text_raw \<open>
173
+ − 1426
\begin{figure}[t]
177
+ − 1427
\begin{minipage}{\textwidth}
565
+ − 1428
\begin{isabelle}\<close>
+ − 1429
ML %grayML\<open>fun print_ss ctxt ss =
157
+ − 1430
let
458
+ − 1431
val {simps, congs, procs, ...} = Raw_Simplifier.dest_ss ss
157
+ − 1432
544
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1433
fun name_sthm (nm, thm) =
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1434
Pretty.enclose (nm ^ ": ") "" [pretty_thm_no_vars ctxt thm]
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1435
fun name_cthm ((_, nm), thm) =
440
+ − 1436
Pretty.enclose (nm ^ ": ") "" [pretty_thm_no_vars ctxt thm]
562
+ − 1437
fun name_trm (nm, trm) =
+ − 1438
Pretty.enclose (nm ^ ": ") "" [pretty_terms ctxt trm]
440
+ − 1439
544
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1440
val pps = [Pretty.big_list "Simplification rules:" (map name_sthm simps),
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1441
Pretty.big_list "Congruences rules:" (map name_cthm congs),
562
+ − 1442
Pretty.big_list "Simproc patterns:" (map name_trm procs)]
157
+ − 1443
in
440
+ − 1444
pps |> Pretty.chunks
+ − 1445
|> pwriteln
565
+ − 1446
end\<close>
+ − 1447
text_raw \<open>
157
+ − 1448
\end{isabelle}
177
+ − 1449
\end{minipage}
458
+ − 1450
\caption{The function @{ML_ind dest_ss in Raw_Simplifier} returns a record containing
163
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1451
all printable information stored in a simpset. We are here only interested in the
231
+ − 1452
simplification rules, congruence rules and simprocs.\label{fig:printss}}
565
+ − 1453
\end{figure}\<close>
+ − 1454
+ − 1455
+ − 1456
text \<open>
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1457
To see how they work, consider the function in Figure~\ref{fig:printss}, which
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1458
prints out some parts of a simpset. If you use it to print out the components of the
458
+ − 1459
empty simpset, i.e., @{ML_ind empty_ss in Raw_Simplifier}
157
+ − 1460
+ − 1461
@{ML_response_fake [display,gray]
163
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1462
"print_ss @{context} empty_ss"
157
+ − 1463
"Simplification rules:
163
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1464
Congruences rules:
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1465
Simproc patterns:"}
157
+ − 1466
+ − 1467
you can see it contains nothing. This simpset is usually not useful, except as a
+ − 1468
building block to build bigger simpsets. For example you can add to @{ML empty_ss}
+ − 1469
the simplification rule @{thm [source] Diff_Int} as follows:
565
+ − 1470
\<close>
+ − 1471
+ − 1472
ML %grayML\<open>val ss1 = put_simpset empty_ss @{context} addsimps [@{thm Diff_Int} RS @{thm eq_reflection}]\<close>
+ − 1473
+ − 1474
text \<open>
162
+ − 1475
Printing then out the components of the simpset gives:
153
+ − 1476
157
+ − 1477
@{ML_response_fake [display,gray]
544
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1478
"print_ss @{context} (Raw_Simplifier.simpset_of ss1)"
157
+ − 1479
"Simplification rules:
158
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1480
??.unknown: A - B \<inter> C \<equiv> A - B \<union> (A - C)
163
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1481
Congruences rules:
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1482
Simproc patterns:"}
157
+ − 1483
368
+ − 1484
\footnote{\bf FIXME: Why does it print out ??.unknown}
158
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1485
554
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1486
Adding also the congruence rule @{thm [source] strong_INF_cong}
565
+ − 1487
\<close>
+ − 1488
+ − 1489
ML %grayML\<open>val ss2 = ss1 |> Simplifier.add_cong (@{thm strong_INF_cong} RS @{thm eq_reflection})\<close>
+ − 1490
+ − 1491
text \<open>
157
+ − 1492
gives
+ − 1493
+ − 1494
@{ML_response_fake [display,gray]
544
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1495
"print_ss @{context} (Raw_Simplifier.simpset_of ss2)"
157
+ − 1496
"Simplification rules:
158
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1497
??.unknown: A - B \<inter> C \<equiv> A - B \<union> (A - C)
157
+ − 1498
Congruences rules:
554
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1499
Complete_Lattices.Inf_class.INFIMUM:
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1500
\<lbrakk>A1 = B1; \<And>x. x \<in> B1 =simp=> C1 x = D1 x\<rbrakk> \<Longrightarrow> INFIMUM A1 C1 \<equiv> INFIMUM B1 D1
163
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1501
Simproc patterns:"}
157
+ − 1502
+ − 1503
Notice that we had to add these lemmas as meta-equations. The @{ML empty_ss}
370
+ − 1504
expects this form of the simplification and congruence rules. This is
+ − 1505
different, if we use for example the simpset @{ML HOL_basic_ss} (see below),
+ − 1506
where rules are usually added as equation. However, even
162
+ − 1507
when adding these lemmas to @{ML empty_ss} we do not end up with anything useful yet.
368
+ − 1508
In the context of HOL, the first really useful simpset is @{ML_ind
+ − 1509
HOL_basic_ss in Simpdata}. While printing out the components of this simpset
157
+ − 1510
+ − 1511
@{ML_response_fake [display,gray]
163
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1512
"print_ss @{context} HOL_basic_ss"
157
+ − 1513
"Simplification rules:
163
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1514
Congruences rules:
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1515
Simproc patterns:"}
157
+ − 1516
370
+ − 1517
also produces ``nothing'', the printout is somewhat misleading. In fact
162
+ − 1518
the @{ML HOL_basic_ss} is setup so that it can solve goals of the
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1519
form
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1520
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1521
\begin{isabelle}
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1522
@{thm TrueI}, @{thm refl[no_vars]}, @{term "t \<equiv> t"} and @{thm FalseE[no_vars]};
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1523
\end{isabelle}
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1524
162
+ − 1525
and also resolve with assumptions. For example:
565
+ − 1526
\<close>
157
+ − 1527
+ − 1528
lemma
368
+ − 1529
shows "True"
+ − 1530
and "t = t"
+ − 1531
and "t \<equiv> t"
+ − 1532
and "False \<Longrightarrow> Foo"
+ − 1533
and "\<lbrakk>A; B; C\<rbrakk> \<Longrightarrow> A"
565
+ − 1534
apply(tactic \<open>ALLGOALS (simp_tac (put_simpset HOL_basic_ss @{context}))\<close>)
157
+ − 1535
done
+ − 1536
565
+ − 1537
text \<open>
162
+ − 1538
This behaviour is not because of simplification rules, but how the subgoaler, solver
369
+ − 1539
and looper are set up in @{ML HOL_basic_ss}.
368
+ − 1540
+ − 1541
The simpset @{ML_ind HOL_ss} is an extension of @{ML HOL_basic_ss} containing
162
+ − 1542
already many useful simplification and congruence rules for the logical
+ − 1543
connectives in HOL.
157
+ − 1544
+ − 1545
@{ML_response_fake [display,gray]
163
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1546
"print_ss @{context} HOL_ss"
157
+ − 1547
"Simplification rules:
158
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1548
Pure.triv_forall_equality: (\<And>x. PROP V) \<equiv> PROP V
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1549
HOL.the_eq_trivial: THE x. x = y \<equiv> y
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1550
HOL.the_sym_eq_trivial: THE ya. y = ya \<equiv> y
157
+ − 1551
\<dots>
+ − 1552
Congruences rules:
+ − 1553
HOL.simp_implies: \<dots>
158
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1554
\<Longrightarrow> (PROP P =simp=> PROP Q) \<equiv> (PROP P' =simp=> PROP Q')
163
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1555
op -->: \<lbrakk>P \<equiv> P'; P' \<Longrightarrow> Q \<equiv> Q'\<rbrakk> \<Longrightarrow> P \<longrightarrow> Q \<equiv> P' \<longrightarrow> Q'
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1556
Simproc patterns:
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1557
\<dots>"}
157
+ − 1558
368
+ − 1559
\begin{readmore}
+ − 1560
The simplifier for HOL is set up in @{ML_file "HOL/Tools/simpdata.ML"}.
+ − 1561
The simpset @{ML HOL_ss} is implemented in @{ML_file "HOL/HOL.thy"}.
+ − 1562
\end{readmore}
+ − 1563
162
+ − 1564
The simplifier is often used to unfold definitions in a proof. For this the
458
+ − 1565
simplifier implements the tactic @{ML_ind rewrite_goal_tac in Raw_Simplifier}.\footnote{\bf FIXME:
243
+ − 1566
see LocalDefs infrastructure.} Suppose for example the
162
+ − 1567
definition
565
+ − 1568
\<close>
162
+ − 1569
+ − 1570
definition "MyTrue \<equiv> True"
+ − 1571
565
+ − 1572
text \<open>
370
+ − 1573
then we can use this tactic to unfold the definition of this constant.
565
+ − 1574
\<close>
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1575
362
+ − 1576
lemma
+ − 1577
shows "MyTrue \<Longrightarrow> True"
565
+ − 1578
apply(tactic \<open>rewrite_goal_tac @{context} @{thms MyTrue_def} 1\<close>)
+ − 1579
txt\<open>producing the goal state
368
+ − 1580
\begin{isabelle}
162
+ − 1581
@{subgoals [display]}
565
+ − 1582
\end{isabelle}\<close>
162
+ − 1583
(*<*)oops(*>*)
+ − 1584
565
+ − 1585
text \<open>
370
+ − 1586
If you want to unfold definitions in \emph{all} subgoals, not just one,
458
+ − 1587
then use the the tactic @{ML_ind rewrite_goals_tac in Raw_Simplifier}.
565
+ − 1588
\<close>
+ − 1589
+ − 1590
+ − 1591
text_raw \<open>
173
+ − 1592
\begin{figure}[p]
+ − 1593
\begin{boxedminipage}{\textwidth}
565
+ − 1594
\begin{isabelle}\<close>
475
+ − 1595
type_synonym prm = "(nat \<times> nat) list"
157
+ − 1596
consts perm :: "prm \<Rightarrow> 'a \<Rightarrow> 'a" ("_ \<bullet> _" [80,80] 80)
+ − 1597
229
+ − 1598
overloading
+ − 1599
perm_nat \<equiv> "perm :: prm \<Rightarrow> nat \<Rightarrow> nat"
+ − 1600
perm_prod \<equiv> "perm :: prm \<Rightarrow> ('a\<times>'b) \<Rightarrow> ('a\<times>'b)"
+ − 1601
perm_list \<equiv> "perm :: prm \<Rightarrow> 'a list \<Rightarrow> 'a list"
+ − 1602
begin
+ − 1603
+ − 1604
fun swap::"nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+ − 1605
where
+ − 1606
"swap a b c = (if c=a then b else (if c=b then a else c))"
153
+ − 1607
229
+ − 1608
primrec perm_nat
+ − 1609
where
+ − 1610
"perm_nat [] c = c"
+ − 1611
| "perm_nat (ab#pi) c = swap (fst ab) (snd ab) (perm_nat pi c)"
157
+ − 1612
229
+ − 1613
fun perm_prod
+ − 1614
where
+ − 1615
"perm_prod pi (x, y) = (pi\<bullet>x, pi\<bullet>y)"
+ − 1616
+ − 1617
primrec perm_list
+ − 1618
where
+ − 1619
"perm_list pi [] = []"
+ − 1620
| "perm_list pi (x#xs) = (pi\<bullet>x)#(perm_list pi xs)"
+ − 1621
+ − 1622
end
157
+ − 1623
+ − 1624
lemma perm_append[simp]:
551
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1625
fixes c::"nat" and pi\<^sub>1 pi\<^sub>2::"prm"
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1626
shows "((pi\<^sub>1@pi\<^sub>2)\<bullet>c) = (pi\<^sub>1\<bullet>(pi\<^sub>2\<bullet>c))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1627
by (induct pi\<^sub>1) (auto)
157
+ − 1628
229
+ − 1629
lemma perm_bij[simp]:
362
+ − 1630
fixes c d::"nat" and pi::"prm"
+ − 1631
shows "(pi\<bullet>c = pi\<bullet>d) = (c = d)"
157
+ − 1632
by (induct pi) (auto)
153
+ − 1633
157
+ − 1634
lemma perm_rev[simp]:
362
+ − 1635
fixes c::"nat" and pi::"prm"
+ − 1636
shows "pi\<bullet>((rev pi)\<bullet>c) = c"
157
+ − 1637
by (induct pi arbitrary: c) (auto)
+ − 1638
+ − 1639
lemma perm_compose:
551
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1640
fixes c::"nat" and pi\<^sub>1 pi\<^sub>2::"prm"
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1641
shows "pi\<^sub>1\<bullet>(pi\<^sub>2\<bullet>c) = (pi\<^sub>1\<bullet>pi\<^sub>2)\<bullet>(pi\<^sub>1\<bullet>c)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1642
by (induct pi\<^sub>2) (auto)
565
+ − 1643
text_raw \<open>
173
+ − 1644
\end{isabelle}
+ − 1645
\end{boxedminipage}
229
+ − 1646
\caption{A simple theory about permutations over @{typ nat}s. The point is that the
157
+ − 1647
lemma @{thm [source] perm_compose} cannot be directly added to the simplifier, as
+ − 1648
it would cause the simplifier to loop. It can still be used as a simplification
229
+ − 1649
rule if the permutation in the right-hand side is sufficiently protected.
+ − 1650
\label{fig:perms}}
565
+ − 1651
\end{figure}\<close>
+ − 1652
+ − 1653
+ − 1654
text \<open>
162
+ − 1655
The simplifier is often used in order to bring terms into a normal form.
+ − 1656
Unfortunately, often the situation arises that the corresponding
+ − 1657
simplification rules will cause the simplifier to run into an infinite
+ − 1658
loop. Consider for example the simple theory about permutations over natural
+ − 1659
numbers shown in Figure~\ref{fig:perms}. The purpose of the lemmas is to
+ − 1660
push permutations as far inside as possible, where they might disappear by
+ − 1661
Lemma @{thm [source] perm_rev}. However, to fully normalise all instances,
+ − 1662
it would be desirable to add also the lemma @{thm [source] perm_compose} to
+ − 1663
the simplifier for pushing permutations over other permutations. Unfortunately,
+ − 1664
the right-hand side of this lemma is again an instance of the left-hand side
209
+ − 1665
and so causes an infinite loop. There seems to be no easy way to reformulate
162
+ − 1666
this rule and so one ends up with clunky proofs like:
565
+ − 1667
\<close>
153
+ − 1668
157
+ − 1669
lemma
551
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1670
fixes c d::"nat" and pi\<^sub>1 pi\<^sub>2::"prm"
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1671
shows "pi\<^sub>1\<bullet>(c, pi\<^sub>2\<bullet>((rev pi\<^sub>1)\<bullet>d)) = (pi\<^sub>1\<bullet>c, (pi\<^sub>1\<bullet>pi\<^sub>2)\<bullet>d)"
157
+ − 1672
apply(simp)
+ − 1673
apply(rule trans)
+ − 1674
apply(rule perm_compose)
+ − 1675
apply(simp)
+ − 1676
done
153
+ − 1677
565
+ − 1678
text \<open>
162
+ − 1679
It is however possible to create a single simplifier tactic that solves
157
+ − 1680
such proofs. The trick is to introduce an auxiliary constant for permutations
162
+ − 1681
and split the simplification into two phases (below actually three). Let
+ − 1682
assume the auxiliary constant is
565
+ − 1683
\<close>
157
+ − 1684
+ − 1685
definition
+ − 1686
perm_aux :: "prm \<Rightarrow> 'a \<Rightarrow> 'a" ("_ \<bullet>aux _" [80,80] 80)
+ − 1687
where
+ − 1688
"pi \<bullet>aux c \<equiv> pi \<bullet> c"
+ − 1689
565
+ − 1690
text \<open>Now the two lemmas\<close>
157
+ − 1691
+ − 1692
lemma perm_aux_expand:
551
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1693
fixes c::"nat" and pi\<^sub>1 pi\<^sub>2::"prm"
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1694
shows "pi\<^sub>1\<bullet>(pi\<^sub>2\<bullet>c) = pi\<^sub>1 \<bullet>aux (pi\<^sub>2\<bullet>c)"
157
+ − 1695
unfolding perm_aux_def by (rule refl)
+ − 1696
+ − 1697
lemma perm_compose_aux:
551
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1698
fixes c::"nat" and pi\<^sub>1 pi\<^sub>2::"prm"
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1699
shows "pi\<^sub>1\<bullet>(pi\<^sub>2\<bullet>aux c) = (pi\<^sub>1\<bullet>pi\<^sub>2) \<bullet>aux (pi\<^sub>1\<bullet>c)"
157
+ − 1700
unfolding perm_aux_def by (rule perm_compose)
+ − 1701
565
+ − 1702
text \<open>
157
+ − 1703
are simple consequence of the definition and @{thm [source] perm_compose}.
+ − 1704
More importantly, the lemma @{thm [source] perm_compose_aux} can be safely
+ − 1705
added to the simplifier, because now the right-hand side is not anymore an instance
162
+ − 1706
of the left-hand side. In a sense it freezes all redexes of permutation compositions
+ − 1707
after one step. In this way, we can split simplification of permutations
213
+ − 1708
into three phases without the user noticing anything about the auxiliary
231
+ − 1709
constant. We first freeze any instance of permutation compositions in the term using
162
+ − 1710
lemma @{thm [source] "perm_aux_expand"} (Line 9);
231
+ − 1711
then simplify all other permutations including pushing permutations over
162
+ − 1712
other permutations by rule @{thm [source] perm_compose_aux} (Line 10); and
+ − 1713
finally ``unfreeze'' all instances of permutation compositions by unfolding
+ − 1714
the definition of the auxiliary constant.
565
+ − 1715
\<close>
+ − 1716
+ − 1717
ML %linenosgray\<open>fun perm_simp_tac ctxt =
157
+ − 1718
let
+ − 1719
val thms1 = [@{thm perm_aux_expand}]
229
+ − 1720
val thms2 = [@{thm perm_append}, @{thm perm_bij}, @{thm perm_rev},
157
+ − 1721
@{thm perm_compose_aux}] @ @{thms perm_prod.simps} @
+ − 1722
@{thms perm_list.simps} @ @{thms perm_nat.simps}
+ − 1723
val thms3 = [@{thm perm_aux_def}]
544
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1724
val ss = put_simpset HOL_basic_ss ctxt
157
+ − 1725
in
544
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1726
simp_tac (ss addsimps thms1)
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1727
THEN' simp_tac (ss addsimps thms2)
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1728
THEN' simp_tac (ss addsimps thms3)
565
+ − 1729
end\<close>
+ − 1730
+ − 1731
text \<open>
209
+ − 1732
For all three phases we have to build simpsets adding specific lemmas. As is sufficient
232
+ − 1733
for our purposes here, we can add these lemmas to @{ML HOL_basic_ss} in order to obtain
162
+ − 1734
the desired results. Now we can solve the following lemma
565
+ − 1735
\<close>
157
+ − 1736
+ − 1737
lemma
551
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1738
fixes c d::"nat" and pi\<^sub>1 pi\<^sub>2::"prm"
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1739
shows "pi\<^sub>1\<bullet>(c, pi\<^sub>2\<bullet>((rev pi\<^sub>1)\<bullet>d)) = (pi\<^sub>1\<bullet>c, (pi\<^sub>1\<bullet>pi\<^sub>2)\<bullet>d)"
565
+ − 1740
apply(tactic \<open>perm_simp_tac @{context} 1\<close>)
157
+ − 1741
done
+ − 1742
152
+ − 1743
565
+ − 1744
text \<open>
209
+ − 1745
in one step. This tactic can deal with most instances of normalising permutations.
+ − 1746
In order to solve all cases we have to deal with corner-cases such as the
162
+ − 1747
lemma being an exact instance of the permutation composition lemma. This can
+ − 1748
often be done easier by implementing a simproc or a conversion. Both will be
+ − 1749
explained in the next two chapters.
+ − 1750
562
+ − 1751
(FIXME: Is it interesting to say something about @{term "(=simp=>)"}?)
157
+ − 1752
+ − 1753
(FIXME: What are the second components of the congruence rules---something to
+ − 1754
do with weak congruence constants?)
+ − 1755
240
+ − 1756
(FIXME: what are @{ML mksimps_pairs}; used in Nominal.thy)
+ − 1757
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1758
(FIXME: explain @{ML simplify} and @{ML "Simplifier.rewrite_rule"} etc.)
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1759
565
+ − 1760
\<close>
+ − 1761
+ − 1762
section \<open>Simprocs\<close>
+ − 1763
+ − 1764
text \<open>
129
+ − 1765
In Isabelle you can also implement custom simplification procedures, called
149
+ − 1766
\emph{simprocs}. Simprocs can be triggered by the simplifier on a specified
+ − 1767
term-pattern and rewrite a term according to a theorem. They are useful in
+ − 1768
cases where a rewriting rule must be produced on ``demand'' or when
+ − 1769
rewriting by simplification is too unpredictable and potentially loops.
129
+ − 1770
+ − 1771
To see how simprocs work, let us first write a simproc that just prints out
132
+ − 1772
the pattern which triggers it and otherwise does nothing. For this
129
+ − 1773
you can use the function:
565
+ − 1774
\<close>
+ − 1775
+ − 1776
ML %linenosgray\<open>fun fail_simproc ctxt redex =
129
+ − 1777
let
543
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1778
val _ = pwriteln (Pretty.block
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1779
[Pretty.str "The redex: ", pretty_cterm ctxt redex])
129
+ − 1780
in
+ − 1781
NONE
565
+ − 1782
end\<close>
+ − 1783
+ − 1784
text \<open>
129
+ − 1785
This function takes a simpset and a redex (a @{ML_type cterm}) as
132
+ − 1786
arguments. In Lines 3 and~4, we first extract the context from the given
129
+ − 1787
simpset and then print out a message containing the redex. The function
+ − 1788
returns @{ML NONE} (standing for an optional @{ML_type thm}) since at the
+ − 1789
moment we are \emph{not} interested in actually rewriting anything. We want
130
+ − 1790
that the simproc is triggered by the pattern @{term "Suc n"}. This can be
149
+ − 1791
done by adding the simproc to the current simpset as follows
565
+ − 1792
\<close>
+ − 1793
+ − 1794
simproc_setup %gray fail ("Suc n") = \<open>K fail_simproc\<close>
+ − 1795
+ − 1796
text \<open>
129
+ − 1797
where the second argument specifies the pattern and the right-hand side
232
+ − 1798
contains the code of the simproc (we have to use @{ML K} since we are ignoring
230
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1799
an argument about morphisms.
130
+ − 1800
After this, the simplifier is aware of the simproc and you can test whether
131
+ − 1801
it fires on the lemma:
565
+ − 1802
\<close>
120
+ − 1803
362
+ − 1804
lemma
+ − 1805
shows "Suc 0 = 1"
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1806
apply(simp)
565
+ − 1807
txt\<open>
378
+ − 1808
\begin{minipage}{\textwidth}
565
+ − 1809
\small\<open>> The redex: Suc 0\<close>\\
+ − 1810
\<open>> The redex: Suc 0\<close>\\
+ − 1811
\end{minipage}\<close>(*<*)oops(*>*)
+ − 1812
+ − 1813
text \<open>
129
+ − 1814
This will print out the message twice: once for the left-hand side and
130
+ − 1815
once for the right-hand side. The reason is that during simplification the
+ − 1816
simplifier will at some point reduce the term @{term "1::nat"} to @{term "Suc
129
+ − 1817
0"}, and then the simproc ``fires'' also on that term.
+ − 1818
131
+ − 1819
We can add or delete the simproc from the current simpset by the usual
132
+ − 1820
\isacommand{declare}-statement. For example the simproc will be deleted
+ − 1821
with the declaration
565
+ − 1822
\<close>
129
+ − 1823
243
+ − 1824
declare [[simproc del: fail]]
129
+ − 1825
565
+ − 1826
text \<open>
129
+ − 1827
If you want to see what happens with just \emph{this} simproc, without any
565
+ − 1828
interference from other rewrite rules, you can call \<open>fail\<close>
129
+ − 1829
as follows:
565
+ − 1830
\<close>
129
+ − 1831
362
+ − 1832
lemma
+ − 1833
shows "Suc 0 = 1"
565
+ − 1834
apply(tactic \<open>simp_tac (put_simpset
+ − 1835
HOL_basic_ss @{context} addsimprocs [@{simproc fail}]) 1\<close>)
129
+ − 1836
(*<*)oops(*>*)
+ − 1837
565
+ − 1838
text \<open>
131
+ − 1839
Now the message shows up only once since the term @{term "1::nat"} is
+ − 1840
left unchanged.
129
+ − 1841
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1842
Setting up a simproc using the command \isacommand{simproc\_setup} will
129
+ − 1843
always add automatically the simproc to the current simpset. If you do not
+ − 1844
want this, then you have to use a slightly different method for setting
243
+ − 1845
up the simproc. First the function @{ML fail_simproc} needs to be modified
129
+ − 1846
to
565
+ − 1847
\<close>
+ − 1848
+ − 1849
ML %grayML\<open>fun fail_simproc' _ ctxt redex =
129
+ − 1850
let
543
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1851
val _ = pwriteln (Pretty.block
562
+ − 1852
[Pretty.str "The redex:", pretty_cterm ctxt redex])
129
+ − 1853
in
+ − 1854
NONE
565
+ − 1855
end\<close>
+ − 1856
+ − 1857
text \<open>
130
+ − 1858
Here the redex is given as a @{ML_type term}, instead of a @{ML_type cterm}
441
+ − 1859
(therefore we printing it out using the function @{ML pretty_term in Pretty}).
149
+ − 1860
We can turn this function into a proper simproc using the function
562
+ − 1861
@{ML Simplifier.make_simproc}:
565
+ − 1862
\<close>
+ − 1863
+ − 1864
ML %grayML\<open>val fail' =
562
+ − 1865
Simplifier.make_simproc @{context} "fail_simproc'"
565
+ − 1866
{lhss = [@{term "Suc n"}], proc = fail_simproc'}\<close>
+ − 1867
+ − 1868
text \<open>
129
+ − 1869
Here the pattern is given as @{ML_type term} (instead of @{ML_type cterm}).
130
+ − 1870
The function also takes a list of patterns that can trigger the simproc.
132
+ − 1871
Now the simproc is set up and can be explicitly added using
458
+ − 1872
@{ML_ind addsimprocs in Raw_Simplifier} to a simpset whenever
132
+ − 1873
needed.
+ − 1874
+ − 1875
Simprocs are applied from inside to outside and from left to right. You can
+ − 1876
see this in the proof
565
+ − 1877
\<close>
129
+ − 1878
362
+ − 1879
lemma
+ − 1880
shows "Suc (Suc 0) = (Suc 1)"
565
+ − 1881
apply(tactic \<open>simp_tac ((put_simpset HOL_basic_ss @{context}) addsimprocs [fail']) 1\<close>)
129
+ − 1882
(*<*)oops(*>*)
+ − 1883
565
+ − 1884
text \<open>
243
+ − 1885
The simproc @{ML fail'} prints out the sequence
129
+ − 1886
130
+ − 1887
@{text [display]
+ − 1888
"> Suc 0
+ − 1889
> Suc (Suc 0)
+ − 1890
> Suc 1"}
+ − 1891
131
+ − 1892
To see how a simproc applies a theorem, let us implement a simproc that
130
+ − 1893
rewrites terms according to the equation:
565
+ − 1894
\<close>
129
+ − 1895
+ − 1896
lemma plus_one:
+ − 1897
shows "Suc n \<equiv> n + 1" by simp
+ − 1898
565
+ − 1899
text \<open>
130
+ − 1900
Simprocs expect that the given equation is a meta-equation, however the
131
+ − 1901
equation can contain preconditions (the simproc then will only fire if the
132
+ − 1902
preconditions can be solved). To see that one has relatively precise control over
131
+ − 1903
the rewriting with simprocs, let us further assume we want that the simproc
+ − 1904
only rewrites terms ``greater'' than @{term "Suc 0"}. For this we can write
565
+ − 1905
\<close>
+ − 1906
+ − 1907
+ − 1908
ML %grayML\<open>fun plus_one_simproc _ ctxt redex =
562
+ − 1909
case Thm.term_of redex of
129
+ − 1910
@{term "Suc 0"} => NONE
565
+ − 1911
| _ => SOME @{thm plus_one}\<close>
+ − 1912
+ − 1913
text \<open>
129
+ − 1914
and set up the simproc as follows.
565
+ − 1915
\<close>
+ − 1916
+ − 1917
ML %grayML\<open>val plus_one =
562
+ − 1918
Simplifier.make_simproc @{context} "sproc +1"
565
+ − 1919
{lhss = [@{term "Suc n"}], proc = plus_one_simproc}\<close>
+ − 1920
+ − 1921
text \<open>
132
+ − 1922
Now the simproc is set up so that it is triggered by terms
130
+ − 1923
of the form @{term "Suc n"}, but inside the simproc we only produce
+ − 1924
a theorem if the term is not @{term "Suc 0"}. The result you can see
131
+ − 1925
in the following proof
565
+ − 1926
\<close>
129
+ − 1927
362
+ − 1928
lemma
+ − 1929
shows "P (Suc (Suc (Suc 0))) (Suc 0)"
565
+ − 1930
apply(tactic \<open>simp_tac (put_simpset HOL_basic_ss @{context} addsimprocs [plus_one]) 1\<close>)
+ − 1931
txt\<open>
131
+ − 1932
where the simproc produces the goal state
368
+ − 1933
\begin{isabelle}
129
+ − 1934
@{subgoals[display]}
368
+ − 1935
\end{isabelle}
565
+ − 1936
\<close>
129
+ − 1937
(*<*)oops(*>*)
+ − 1938
565
+ − 1939
text \<open>
133
+ − 1940
As usual with rewriting you have to worry about looping: you already have
243
+ − 1941
a loop with @{ML plus_one}, if you apply it with the default simpset (because
+ − 1942
the default simpset contains a rule which just does the opposite of @{ML plus_one},
132
+ − 1943
namely rewriting @{text [quotes] "+ 1"} to a successor). So you have to be careful
+ − 1944
in choosing the right simpset to which you add a simproc.
130
+ − 1945
132
+ − 1946
Next let us implement a simproc that replaces terms of the form @{term "Suc n"}
565
+ − 1947
with the number \<open>n\<close> increased by one. First we implement a function that
132
+ − 1948
takes a term and produces the corresponding integer value.
565
+ − 1949
\<close>
+ − 1950
+ − 1951
ML %grayML\<open>fun dest_suc_trm ((Const (@{const_name "Suc"}, _)) $ t) = 1 + dest_suc_trm t
+ − 1952
| dest_suc_trm t = snd (HOLogic.dest_number t)\<close>
+ − 1953
+ − 1954
text \<open>
316
+ − 1955
It uses the library function @{ML_ind dest_number in HOLogic} that transforms
130
+ − 1956
(Isabelle) terms, like @{term "0::nat"}, @{term "1::nat"}, @{term "2::nat"} and so
131
+ − 1957
on, into integer values. This function raises the exception @{ML TERM}, if
130
+ − 1958
the term is not a number. The next function expects a pair consisting of a term
565
+ − 1959
\<open>t\<close> (containing @{term Suc}s) and the corresponding integer value \<open>n\<close>.
+ − 1960
\<close>
+ − 1961
+ − 1962
ML %linenosgray\<open>fun get_thm ctxt (t, n) =
130
+ − 1963
let
+ − 1964
val num = HOLogic.mk_number @{typ "nat"} n
132
+ − 1965
val goal = Logic.mk_equals (t, num)
130
+ − 1966
in
214
+ − 1967
Goal.prove ctxt [] [] goal (K (Arith_Data.arith_tac ctxt 1))
565
+ − 1968
end\<close>
+ − 1969
+ − 1970
text \<open>
132
+ − 1971
From the integer value it generates the corresponding number term, called
565
+ − 1972
\<open>num\<close> (Line 3), and then generates the meta-equation \<open>t \<equiv> num\<close>
132
+ − 1973
(Line 4), which it proves by the arithmetic tactic in Line 6.
+ − 1974
219
+ − 1975
For our purpose at the moment, proving the meta-equation using @{ML
+ − 1976
arith_tac in Arith_Data} is fine, but there is also an alternative employing
+ − 1977
the simplifier with a special simpset. For the kind of lemmas we
565
+ − 1978
want to prove here, the simpset \<open>num_ss\<close> should suffice.
+ − 1979
\<close>
+ − 1980
+ − 1981
ML %grayML\<open>fun get_thm_alt ctxt (t, n) =
132
+ − 1982
let
+ − 1983
val num = HOLogic.mk_number @{typ "nat"} n
+ − 1984
val goal = Logic.mk_equals (t, num)
544
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1985
val num_ss = put_simpset HOL_ss ctxt addsimps @{thms semiring_norm}
132
+ − 1986
in
+ − 1987
Goal.prove ctxt [] [] goal (K (simp_tac num_ss 1))
565
+ − 1988
end\<close>
+ − 1989
+ − 1990
text \<open>
132
+ − 1991
The advantage of @{ML get_thm_alt} is that it leaves very little room for
+ − 1992
something to go wrong; in contrast it is much more difficult to predict
219
+ − 1993
what happens with @{ML arith_tac in Arith_Data}, especially in more complicated
231
+ − 1994
circumstances. The disadvantage of @{ML get_thm_alt} is to find a simpset
132
+ − 1995
that is sufficiently powerful to solve every instance of the lemmas
+ − 1996
we like to prove. This requires careful tuning, but is often necessary in
+ − 1997
``production code''.\footnote{It would be of great help if there is another
+ − 1998
way than tracing the simplifier to obtain the lemmas that are successfully
+ − 1999
applied during simplification. Alas, there is none.}
+ − 2000
+ − 2001
Anyway, either version can be used in the function that produces the actual
+ − 2002
theorem for the simproc.
565
+ − 2003
\<close>
+ − 2004
+ − 2005
ML %grayML\<open>fun nat_number_simproc _ ctxt ct =
562
+ − 2006
SOME (get_thm_alt ctxt (Thm.term_of ct, dest_suc_trm (Thm.term_of ct)))
565
+ − 2007
handle TERM _ => NONE\<close>
+ − 2008
+ − 2009
text \<open>
243
+ − 2010
This function uses the fact that @{ML dest_suc_trm} might raise an exception
130
+ − 2011
@{ML TERM}. In this case there is nothing that can be rewritten and therefore no
131
+ − 2012
theorem is produced (i.e.~the function returns @{ML NONE}). To try out the simproc
+ − 2013
on an example, you can set it up as follows:
565
+ − 2014
\<close>
+ − 2015
+ − 2016
ML %grayML\<open>val nat_number =
562
+ − 2017
Simplifier.make_simproc @{context} "nat_number"
565
+ − 2018
{lhss = [@{term "Suc n"}], proc = nat_number_simproc}\<close>
+ − 2019
+ − 2020
+ − 2021
text \<open>
130
+ − 2022
Now in the lemma
565
+ − 2023
\<close>
129
+ − 2024
362
+ − 2025
lemma
+ − 2026
shows "P (Suc (Suc 2)) (Suc 99) (0::nat) (Suc 4 + Suc 0) (Suc (0 + 0))"
565
+ − 2027
apply(tactic \<open>simp_tac (put_simpset HOL_ss @{context} addsimprocs [nat_number]) 1\<close>)
+ − 2028
txt \<open>
130
+ − 2029
you obtain the more legible goal state
368
+ − 2030
\begin{isabelle}
129
+ − 2031
@{subgoals [display]}
565
+ − 2032
\end{isabelle}\<close>
129
+ − 2033
(*<*)oops(*>*)
+ − 2034
565
+ − 2035
text \<open>
132
+ − 2036
where the simproc rewrites all @{term "Suc"}s except in the last argument. There it cannot
130
+ − 2037
rewrite anything, because it does not know how to transform the term @{term "Suc (0 + 0)"}
+ − 2038
into a number. To solve this problem have a look at the next exercise.
+ − 2039
+ − 2040
\begin{exercise}\label{ex:addsimproc}
551
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 2041
Write a simproc that replaces terms of the form @{term "t\<^sub>1 + t\<^sub>2"} by their
130
+ − 2042
result. You can assume the terms are ``proper'' numbers, that is of the form
+ − 2043
@{term "0::nat"}, @{term "1::nat"}, @{term "2::nat"} and so on.
+ − 2044
\end{exercise}
+ − 2045
+ − 2046
(FIXME: We did not do anything with morphisms. Anything interesting
+ − 2047
one can say about them?)
565
+ − 2048
\<close>
+ − 2049
+ − 2050
section \<open>Conversions\label{sec:conversion}\<close>
+ − 2051
+ − 2052
text \<open>
406
+ − 2053
Conversions are a thin layer on top of Isabelle's inference kernel, and can
+ − 2054
be viewed as a controllable, bare-bone version of Isabelle's simplifier.
412
+ − 2055
The purpose of conversions is to manipulate @{ML_type cterm}s. However,
406
+ − 2056
we will also show in this section how conversions can be applied to theorems
412
+ − 2057
and to goal states. The type of conversions is
565
+ − 2058
\<close>
+ − 2059
+ − 2060
ML %grayML\<open>type conv = cterm -> thm\<close>
+ − 2061
+ − 2062
text \<open>
147
+ − 2063
whereby the produced theorem is always a meta-equality. A simple conversion
316
+ − 2064
is the function @{ML_ind all_conv in Conv}, which maps a @{ML_type cterm} to an
147
+ − 2065
instance of the (meta)reflexivity theorem. For example:
135
+ − 2066
145
+ − 2067
@{ML_response_fake [display,gray]
146
+ − 2068
"Conv.all_conv @{cterm \"Foo \<or> Bar\"}"
+ − 2069
"Foo \<or> Bar \<equiv> Foo \<or> Bar"}
+ − 2070
316
+ − 2071
Another simple conversion is @{ML_ind no_conv in Conv} which always raises the
147
+ − 2072
exception @{ML CTERM}.
135
+ − 2073
145
+ − 2074
@{ML_response_fake [display,gray]
+ − 2075
"Conv.no_conv @{cterm True}"
+ − 2076
"*** Exception- CTERM (\"no conversion\", []) raised"}
+ − 2077
316
+ − 2078
A more interesting conversion is the function @{ML_ind beta_conversion in Thm}: it
160
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2079
produces a meta-equation between a term and its beta-normal form. For example
142
+ − 2080
145
+ − 2081
@{ML_response_fake [display,gray]
146
+ − 2082
"let
+ − 2083
val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
+ − 2084
val two = @{cterm \"2::nat\"}
+ − 2085
val ten = @{cterm \"10::nat\"}
513
+ − 2086
val ctrm = Thm.apply (Thm.apply add two) ten
146
+ − 2087
in
291
+ − 2088
Thm.beta_conversion true ctrm
146
+ − 2089
end"
+ − 2090
"((\<lambda>x y. x + y) 2) 10 \<equiv> 2 + 10"}
+ − 2091
291
+ − 2092
If you run this example, you will notice that the actual response is the
+ − 2093
seemingly nonsensical @{term
+ − 2094
"2 + 10 \<equiv> 2 + (10::nat)"}. The reason is that the pretty-printer for
405
+ − 2095
@{ML_type cterm}s eta-normalises (sic) terms and therefore produces this output.
291
+ − 2096
If we get hold of the ``raw'' representation of the produced theorem,
+ − 2097
we obtain the expected result.
+ − 2098
147
+ − 2099
@{ML_response [display,gray]
+ − 2100
"let
+ − 2101
val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
+ − 2102
val two = @{cterm \"2::nat\"}
+ − 2103
val ten = @{cterm \"10::nat\"}
513
+ − 2104
val ctrm = Thm.apply (Thm.apply add two) ten
147
+ − 2105
in
291
+ − 2106
Thm.prop_of (Thm.beta_conversion true ctrm)
147
+ − 2107
end"
566
6103b0eadbf2
tuned parser for patterns in ML_response... antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
diff
changeset
+ − 2108
"Const (\"Pure.eq\",_) $
6103b0eadbf2
tuned parser for patterns in ML_response... antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
diff
changeset
+ − 2109
(Abs (\"x\",_,Abs (\"y\",_,_)) $_$_) $
6103b0eadbf2
tuned parser for patterns in ML_response... antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
diff
changeset
+ − 2110
(Const (\"Groups.plus_class.plus\",_) $ _ $ _)"}
142
+ − 2111
512
+ − 2112
or in the pretty-printed form
509
+ − 2113
+ − 2114
@{ML_response_fake [display,gray]
+ − 2115
"let
512
+ − 2116
val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
+ − 2117
val two = @{cterm \"2::nat\"}
+ − 2118
val ten = @{cterm \"10::nat\"}
513
+ − 2119
val ctrm = Thm.apply (Thm.apply add two) ten
512
+ − 2120
val ctxt = @{context}
+ − 2121
|> Config.put eta_contract false
+ − 2122
|> Config.put show_abbrevs false
509
+ − 2123
in
+ − 2124
Thm.prop_of (Thm.beta_conversion true ctrm)
+ − 2125
|> pretty_term ctxt
+ − 2126
|> pwriteln
+ − 2127
end"
+ − 2128
"(\<lambda>x y. x + y) 2 10 \<equiv> 2 + 10"}
+ − 2129
291
+ − 2130
The argument @{ML true} in @{ML beta_conversion in Thm} indicates that
243
+ − 2131
the right-hand side should be fully beta-normalised. If instead
147
+ − 2132
@{ML false} is given, then only a single beta-reduction is performed
291
+ − 2133
on the outer-most level.
146
+ − 2134
147
+ − 2135
The main point of conversions is that they can be used for rewriting
291
+ − 2136
@{ML_type cterm}s. One example is the function
316
+ − 2137
@{ML_ind rewr_conv in Conv}, which expects a meta-equation as an
291
+ − 2138
argument. Suppose the following meta-equation.
+ − 2139
565
+ − 2140
\<close>
135
+ − 2141
362
+ − 2142
lemma true_conj1:
+ − 2143
shows "True \<and> P \<equiv> P" by simp
135
+ − 2144
565
+ − 2145
text \<open>
291
+ − 2146
It can be used for example to rewrite @{term "True \<and> (Foo \<longrightarrow> Bar)"}
+ − 2147
to @{term "Foo \<longrightarrow> Bar"}. The code is as follows.
139
+ − 2148
145
+ − 2149
@{ML_response_fake [display,gray]
146
+ − 2150
"let
149
+ − 2151
val ctrm = @{cterm \"True \<and> (Foo \<longrightarrow> Bar)\"}
146
+ − 2152
in
+ − 2153
Conv.rewr_conv @{thm true_conj1} ctrm
+ − 2154
end"
+ − 2155
"True \<and> (Foo \<longrightarrow> Bar) \<equiv> Foo \<longrightarrow> Bar"}
139
+ − 2156
316
+ − 2157
Note, however, that the function @{ML_ind rewr_conv in Conv} only rewrites the
160
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2158
outer-most level of the @{ML_type cterm}. If the given @{ML_type cterm} does not match
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2159
exactly the
316
+ − 2160
left-hand side of the theorem, then @{ML_ind rewr_conv in Conv} fails, raising
147
+ − 2161
the exception @{ML CTERM}.
146
+ − 2162
+ − 2163
This very primitive way of rewriting can be made more powerful by
+ − 2164
combining several conversions into one. For this you can use conversion
369
+ − 2165
combinators. The simplest conversion combinator is @{ML_ind then_conv in Conv},
146
+ − 2166
which applies one conversion after another. For example
139
+ − 2167
145
+ − 2168
@{ML_response_fake [display,gray]
146
+ − 2169
"let
147
+ − 2170
val conv1 = Thm.beta_conversion false
146
+ − 2171
val conv2 = Conv.rewr_conv @{thm true_conj1}
513
+ − 2172
val ctrm = Thm.apply @{cterm \"\<lambda>x. x \<and> False\"} @{cterm \"True\"}
146
+ − 2173
in
+ − 2174
(conv1 then_conv conv2) ctrm
+ − 2175
end"
145
+ − 2176
"(\<lambda>x. x \<and> False) True \<equiv> False"}
139
+ − 2177
147
+ − 2178
where we first beta-reduce the term and then rewrite according to
291
+ − 2179
@{thm [source] true_conj1}. (When running this example recall the
+ − 2180
problem with the pretty-printer normalising all terms.)
147
+ − 2181
369
+ − 2182
The conversion combinator @{ML_ind else_conv in Conv} tries out the
146
+ − 2183
first one, and if it does not apply, tries the second. For example
+ − 2184
145
+ − 2185
@{ML_response_fake [display,gray]
146
+ − 2186
"let
147
+ − 2187
val conv = Conv.rewr_conv @{thm true_conj1} else_conv Conv.all_conv
146
+ − 2188
val ctrm1 = @{cterm \"True \<and> Q\"}
+ − 2189
val ctrm2 = @{cterm \"P \<or> (True \<and> Q)\"}
+ − 2190
in
+ − 2191
(conv ctrm1, conv ctrm2)
+ − 2192
end"
147
+ − 2193
"(True \<and> Q \<equiv> Q, P \<or> True \<and> Q \<equiv> P \<or> True \<and> Q)"}
146
+ − 2194
406
+ − 2195
Here the conversion @{thm [source] true_conj1} only applies
146
+ − 2196
in the first case, but fails in the second. The whole conversion
256
+ − 2197
does not fail, however, because the combinator @{ML else_conv in Conv} will then
405
+ − 2198
try out @{ML all_conv in Conv}, which always succeeds. The same
+ − 2199
behaviour can also be achieved with conversion combinator @{ML_ind try_conv in Conv}.
174
+ − 2200
For example
+ − 2201
+ − 2202
@{ML_response_fake [display,gray]
291
+ − 2203
"let
+ − 2204
val conv = Conv.try_conv (Conv.rewr_conv @{thm true_conj1})
+ − 2205
val ctrm = @{cterm \"True \<or> P\"}
+ − 2206
in
+ − 2207
conv ctrm
+ − 2208
end"
174
+ − 2209
"True \<or> P \<equiv> True \<or> P"}
+ − 2210
405
+ − 2211
Rewriting with more than one theorem can be done using the function
424
+ − 2212
@{ML_ind rewrs_conv in Conv} from the structure @{ML_struct Conv}.
405
+ − 2213
+ − 2214
149
+ − 2215
Apart from the function @{ML beta_conversion in Thm}, which is able to fully
+ − 2216
beta-normalise a term, the conversions so far are restricted in that they
147
+ − 2217
only apply to the outer-most level of a @{ML_type cterm}. In what follows we
369
+ − 2218
will lift this restriction. The combinators @{ML_ind fun_conv in Conv}
+ − 2219
and @{ML_ind arg_conv in Conv} will apply
291
+ − 2220
a conversion to the first, respectively second, argument of an application.
+ − 2221
For example
139
+ − 2222
145
+ − 2223
@{ML_response_fake [display,gray]
146
+ − 2224
"let
291
+ − 2225
val conv = Conv.arg_conv (Conv.rewr_conv @{thm true_conj1})
146
+ − 2226
val ctrm = @{cterm \"P \<or> (True \<and> Q)\"}
+ − 2227
in
291
+ − 2228
conv ctrm
146
+ − 2229
end"
+ − 2230
"P \<or> (True \<and> Q) \<equiv> P \<or> Q"}
139
+ − 2231
565
+ − 2232
The reason for this behaviour is that \<open>(op \<or>)\<close> expects two
+ − 2233
arguments. Therefore the term must be of the form \<open>(Const \<dots> $ t1) $ t2\<close>. The
+ − 2234
conversion is then applied to \<open>t2\<close>, which in the example above
291
+ − 2235
stands for @{term "True \<and> Q"}. The function @{ML fun_conv in Conv} would apply
565
+ − 2236
the conversion to the term \<open>(Const \<dots> $ t1)\<close>.
147
+ − 2237
316
+ − 2238
The function @{ML_ind abs_conv in Conv} applies a conversion under an
291
+ − 2239
abstraction. For example:
139
+ − 2240
147
+ − 2241
@{ML_response_fake [display,gray]
+ − 2242
"let
243
+ − 2243
val conv = Conv.rewr_conv @{thm true_conj1}
291
+ − 2244
val ctrm = @{cterm \"\<lambda>P. True \<and> (P \<and> Foo)\"}
147
+ − 2245
in
243
+ − 2246
Conv.abs_conv (K conv) @{context} ctrm
147
+ − 2247
end"
291
+ − 2248
"\<lambda>P. True \<and> (P \<and> Foo) \<equiv> \<lambda>P. P \<and> Foo"}
147
+ − 2249
291
+ − 2250
Note that this conversion needs a context as an argument. We also give the
565
+ − 2251
conversion as \<open>(K conv)\<close>, which is a function that ignores its
291
+ − 2252
argument (the argument being a sufficiently freshened version of the
+ − 2253
variable that is abstracted and a context). The conversion that goes under
316
+ − 2254
an application is @{ML_ind combination_conv in Conv}. It expects two
291
+ − 2255
conversions as arguments, each of which is applied to the corresponding
292
+ − 2256
``branch'' of the application. An abbreviation for this conversion is the
316
+ − 2257
function @{ML_ind comb_conv in Conv}, which applies the same conversion
292
+ − 2258
to both branches.
147
+ − 2259
160
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2260
We can now apply all these functions in a conversion that recursively
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2261
descends a term and applies a ``@{thm [source] true_conj1}''-conversion
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2262
in all possible positions.
565
+ − 2263
\<close>
+ − 2264
+ − 2265
ML %linenosgray\<open>fun true_conj1_conv ctxt ctrm =
147
+ − 2266
case (Thm.term_of ctrm) of
562
+ − 2267
@{term "(\<and>)"} $ @{term True} $ _ =>
405
+ − 2268
(Conv.arg_conv (true_conj1_conv ctxt) then_conv
147
+ − 2269
Conv.rewr_conv @{thm true_conj1}) ctrm
405
+ − 2270
| _ $ _ => Conv.comb_conv (true_conj1_conv ctxt) ctrm
+ − 2271
| Abs _ => Conv.abs_conv (fn (_, ctxt) => true_conj1_conv ctxt) ctxt ctrm
565
+ − 2272
| _ => Conv.all_conv ctrm\<close>
+ − 2273
+ − 2274
text \<open>
+ − 2275
This function ``fires'' if the term is of the form \<open>(True \<and> \<dots>)\<close>.
406
+ − 2276
It descends under applications (Line 6) and abstractions
+ − 2277
(Line 7); otherwise it leaves the term unchanged (Line 8). In Line 2
149
+ − 2278
we need to transform the @{ML_type cterm} into a @{ML_type term} in order
+ − 2279
to be able to pattern-match the term. To see this
160
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2280
conversion in action, consider the following example:
139
+ − 2281
147
+ − 2282
@{ML_response_fake [display,gray]
+ − 2283
"let
405
+ − 2284
val conv = true_conj1_conv @{context}
+ − 2285
val ctrm = @{cterm \"distinct [1::nat, x] \<longrightarrow> True \<and> 1 \<noteq> x\"}
147
+ − 2286
in
291
+ − 2287
conv ctrm
147
+ − 2288
end"
145
+ − 2289
"distinct [1, x] \<longrightarrow> True \<and> 1 \<noteq> x \<equiv> distinct [1, x] \<longrightarrow> 1 \<noteq> x"}
565
+ − 2290
\<close>
+ − 2291
+ − 2292
text \<open>
405
+ − 2293
Conversions that traverse terms, like @{ML true_conj1_conv} above, can be
+ − 2294
implemented more succinctly with the combinators @{ML_ind bottom_conv in
424
+ − 2295
Conv} and @{ML_ind top_conv in Conv}. For example:
565
+ − 2296
\<close>
+ − 2297
+ − 2298
ML %grayML\<open>fun true_conj1_conv ctxt =
405
+ − 2299
let
+ − 2300
val conv = Conv.try_conv (Conv.rewr_conv @{thm true_conj1})
+ − 2301
in
424
+ − 2302
Conv.bottom_conv (K conv) ctxt
565
+ − 2303
end\<close>
+ − 2304
+ − 2305
text \<open>
408
+ − 2306
If we regard terms as trees with variables and constants on the top, then
424
+ − 2307
@{ML bottom_conv in Conv} traverses first the the term and
405
+ − 2308
on the ``way down'' applies the conversion, whereas @{ML top_conv in
424
+ − 2309
Conv} applies the conversion on the ``way up''. To see the difference,
412
+ − 2310
assume the following two meta-equations
565
+ − 2311
\<close>
405
+ − 2312
+ − 2313
lemma conj_assoc:
+ − 2314
fixes A B C::bool
+ − 2315
shows "A \<and> (B \<and> C) \<equiv> (A \<and> B) \<and> C"
+ − 2316
and "(A \<and> B) \<and> C \<equiv> A \<and> (B \<and> C)"
+ − 2317
by simp_all
+ − 2318
565
+ − 2319
text \<open>
+ − 2320
and the @{ML_type cterm} \<open>(a \<and> (b \<and> c)) \<and> d\<close>. Here you should
405
+ − 2321
pause for a moment to be convinced that rewriting top-down and bottom-up
412
+ − 2322
according to the two meta-equations produces two results. Below these
+ − 2323
two results are calculated.
405
+ − 2324
+ − 2325
@{ML_response_fake [display, gray]
+ − 2326
"let
+ − 2327
val ctxt = @{context}
424
+ − 2328
val conv = Conv.try_conv (Conv.rewrs_conv @{thms conj_assoc})
+ − 2329
val conv_top = Conv.top_conv (K conv) ctxt
+ − 2330
val conv_bot = Conv.bottom_conv (K conv) ctxt
405
+ − 2331
val ctrm = @{cterm \"(a \<and> (b \<and> c)) \<and> d\"}
+ − 2332
in
+ − 2333
(conv_top ctrm, conv_bot ctrm)
+ − 2334
end"
+ − 2335
"(\"(a \<and> (b \<and> c)) \<and> d \<equiv> a \<and> (b \<and> (c \<and> d))\",
+ − 2336
\"(a \<and> (b \<and> c)) \<and> d \<equiv> (a \<and> b) \<and> (c \<and> d)\")"}
139
+ − 2337
412
+ − 2338
To see how much control you have over rewriting with conversions, let us
147
+ − 2339
make the task a bit more complicated by rewriting according to the rule
565
+ − 2340
\<open>true_conj1\<close>, but only in the first arguments of @{term If}s. Then
147
+ − 2341
the conversion should be as follows.
565
+ − 2342
\<close>
+ − 2343
+ − 2344
ML %grayML\<open>fun if_true1_simple_conv ctxt ctrm =
147
+ − 2345
case Thm.term_of ctrm of
142
+ − 2346
Const (@{const_name If}, _) $ _ =>
405
+ − 2347
Conv.arg_conv (true_conj1_conv ctxt) ctrm
+ − 2348
| _ => Conv.no_conv ctrm
+ − 2349
565
+ − 2350
val if_true1_conv = Conv.top_sweep_conv if_true1_simple_conv\<close>
+ − 2351
+ − 2352
text \<open>
405
+ − 2353
In the first function we only treat the top-most redex and also only the
412
+ − 2354
success case. As default we return @{ML no_conv in Conv}. To apply this
+ − 2355
``simple'' conversion inside a term, we use in the last line the combinator @{ML_ind
424
+ − 2356
top_sweep_conv in Conv}, which traverses the term starting from the
406
+ − 2357
root and applies it to all the redexes on the way, but stops in each branch as
+ − 2358
soon as it found one redex. Here is an example for this conversion:
405
+ − 2359
139
+ − 2360
145
+ − 2361
@{ML_response_fake [display,gray]
147
+ − 2362
"let
405
+ − 2363
val ctrm = @{cterm \"if P (True \<and> 1 \<noteq> (2::nat))
+ − 2364
then True \<and> True else True \<and> False\"}
147
+ − 2365
in
405
+ − 2366
if_true1_conv @{context} ctrm
147
+ − 2367
end"
+ − 2368
"if P (True \<and> 1 \<noteq> 2) then True \<and> True else True \<and> False
+ − 2369
\<equiv> if P (1 \<noteq> 2) then True \<and> True else True \<and> False"}
565
+ − 2370
\<close>
+ − 2371
+ − 2372
text \<open>
147
+ − 2373
So far we only applied conversions to @{ML_type cterm}s. Conversions can, however,
316
+ − 2374
also work on theorems using the function @{ML_ind fconv_rule in Conv}. As an example,
412
+ − 2375
consider again the conversion @{ML true_conj1_conv} and the lemma:
565
+ − 2376
\<close>
147
+ − 2377
362
+ − 2378
lemma foo_test:
+ − 2379
shows "P \<or> (True \<and> \<not>P)" by simp
147
+ − 2380
565
+ − 2381
text \<open>
412
+ − 2382
Using the conversion you can transform this theorem into a
291
+ − 2383
new theorem as follows
147
+ − 2384
+ − 2385
@{ML_response_fake [display,gray]
291
+ − 2386
"let
405
+ − 2387
val conv = Conv.fconv_rule (true_conj1_conv @{context})
291
+ − 2388
val thm = @{thm foo_test}
+ − 2389
in
+ − 2390
conv thm
+ − 2391
end"
147
+ − 2392
"?P \<or> \<not> ?P"}
+ − 2393
412
+ − 2394
Finally, Isabelle provides function @{ML_ind CONVERSION in Tactical}
+ − 2395
for turning conversions into tactics. This needs some predefined conversion
+ − 2396
combinators that traverse a goal
410
+ − 2397
state and can selectively apply the conversion. The combinators for
+ − 2398
the goal state are:
291
+ − 2399
+ − 2400
\begin{itemize}
369
+ − 2401
\item @{ML_ind params_conv in Conv} for converting under parameters
565
+ − 2402
(i.e.~where a goal state is of the form \<open>\<And>x. P x \<Longrightarrow> Q x\<close>)
412
+ − 2403
+ − 2404
\item @{ML_ind prems_conv in Conv} for applying a conversion to
+ − 2405
premises of a goal state, and
291
+ − 2406
369
+ − 2407
\item @{ML_ind concl_conv in Conv} for applying a conversion to the
412
+ − 2408
conclusion of a goal state.
291
+ − 2409
\end{itemize}
139
+ − 2410
405
+ − 2411
Assume we want to apply @{ML true_conj1_conv} only in the conclusion
160
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2412
of the goal, and @{ML if_true1_conv} should only apply to the premises.
145
+ − 2413
Here is a tactic doing exactly that:
565
+ − 2414
\<close>
+ − 2415
+ − 2416
ML %grayML\<open>fun true1_tac ctxt =
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2417
CONVERSION
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2418
(Conv.params_conv ~1 (fn ctxt =>
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2419
(Conv.prems_conv ~1 (if_true1_conv ctxt) then_conv
565
+ − 2420
Conv.concl_conv ~1 (true_conj1_conv ctxt))) ctxt)\<close>
+ − 2421
+ − 2422
text \<open>
406
+ − 2423
We call the conversions with the argument @{ML "~1"}. By this we
+ − 2424
analyse all parameters, all premises and the conclusion of a goal
+ − 2425
state. If we call @{ML concl_conv in Conv} with
565
+ − 2426
a non-negative number, say \<open>n\<close>, then this conversions will
+ − 2427
skip the first \<open>n\<close> premises and applies the conversion to the
406
+ − 2428
``rest'' (similar for parameters and conclusions). To test the
+ − 2429
tactic, consider the proof
565
+ − 2430
\<close>
139
+ − 2431
142
+ − 2432
lemma
+ − 2433
"if True \<and> P then P else True \<and> False \<Longrightarrow>
148
+ − 2434
(if True \<and> Q then True \<and> Q else P) \<longrightarrow> True \<and> (True \<and> Q)"
565
+ − 2435
apply(tactic \<open>true1_tac @{context} 1\<close>)
+ − 2436
txt \<open>where the tactic yields the goal state
368
+ − 2437
\begin{isabelle}
177
+ − 2438
@{subgoals [display]}
565
+ − 2439
\end{isabelle}\<close>
142
+ − 2440
(*<*)oops(*>*)
135
+ − 2441
565
+ − 2442
text \<open>
148
+ − 2443
As you can see, the premises are rewritten according to @{ML if_true1_conv}, while
410
+ − 2444
the conclusion according to @{ML true_conj1_conv}. If we only have one
+ − 2445
conversion that should be uniformly applied to the whole goal state, we
424
+ − 2446
can simplify @{ML true1_tac} using @{ML_ind top_conv in Conv}.
412
+ − 2447
+ − 2448
Conversions are also be helpful for constructing meta-equality theorems.
332
+ − 2449
Such theorems are often definitions. As an example consider the following
+ − 2450
two ways of defining the identity function in Isabelle.
565
+ − 2451
\<close>
332
+ − 2452
+ − 2453
definition id1::"'a \<Rightarrow> 'a"
+ − 2454
where "id1 x \<equiv> x"
+ − 2455
+ − 2456
definition id2::"'a \<Rightarrow> 'a"
+ − 2457
where "id2 \<equiv> \<lambda>x. x"
+ − 2458
565
+ − 2459
text \<open>
335
+ − 2460
Although both definitions define the same function, the theorems @{thm
412
+ − 2461
[source] id1_def} and @{thm [source] id2_def} are different meta-equations. However it is
+ − 2462
easy to transform one into the other. The function @{ML_ind abs_def
+ − 2463
in Drule} from the structure @{ML_struct Drule} can automatically transform
+ − 2464
theorem @{thm [source] id1_def}
334
+ − 2465
into @{thm [source] id2_def} by abstracting all variables on the
+ − 2466
left-hand side in the right-hand side.
332
+ − 2467
+ − 2468
@{ML_response_fake [display,gray]
+ − 2469
"Drule.abs_def @{thm id1_def}"
+ − 2470
"id1 \<equiv> \<lambda>x. x"}
+ − 2471
406
+ − 2472
Unfortunately, Isabelle has no built-in function that transforms the
+ − 2473
theorems in the other direction. We can implement one using
334
+ − 2474
conversions. The feature of conversions we are using is that if we apply a
+ − 2475
@{ML_type cterm} to a conversion we obtain a meta-equality theorem (recall
+ − 2476
that the type of conversions is an abbreviation for
+ − 2477
@{ML_type "cterm -> thm"}). The code of the transformation is below.
565
+ − 2478
\<close>
+ − 2479
+ − 2480
ML %linenosgray\<open>fun unabs_def ctxt def =
332
+ − 2481
let
562
+ − 2482
val (lhs, rhs) = Thm.dest_equals (Thm.cprop_of def)
+ − 2483
val xs = strip_abs_vars (Thm.term_of rhs)
332
+ − 2484
val (_, ctxt') = Variable.add_fixes (map fst xs) ctxt
+ − 2485
562
+ − 2486
val cxs = map (Thm.cterm_of ctxt' o Free) xs
332
+ − 2487
val new_lhs = Drule.list_comb (lhs, cxs)
+ − 2488
+ − 2489
fun get_conv [] = Conv.rewr_conv def
334
+ − 2490
| get_conv (_::xs) = Conv.fun_conv (get_conv xs)
332
+ − 2491
in
+ − 2492
get_conv xs new_lhs |>
475
+ − 2493
singleton (Proof_Context.export ctxt' ctxt)
565
+ − 2494
end\<close>
+ − 2495
+ − 2496
text \<open>
332
+ − 2497
In Line 3 we destruct the meta-equality into the @{ML_type cterm}s
+ − 2498
corresponding to the left-hand and right-hand side of the meta-equality. The
+ − 2499
assumption in @{ML unabs_def} is that the right-hand side is an
334
+ − 2500
abstraction. We compute the possibly empty list of abstracted variables in
369
+ − 2501
Line 4 using the function @{ML_ind strip_abs_vars in Term}. For this we have to
412
+ − 2502
first transform the @{ML_type cterm} into a @{ML_type term}. In Line 5 we
565
+ − 2503
import these variables into the context \<open>ctxt'\<close>, in order to
334
+ − 2504
export them again in Line 15. In Line 8 we certify the list of variables,
+ − 2505
which in turn we apply to the @{ML_text lhs} of the definition using the
+ − 2506
function @{ML_ind list_comb in Drule}. In Line 11 and 12 we generate the
+ − 2507
conversion according to the length of the list of (abstracted) variables. If
+ − 2508
there are none, then we just return the conversion corresponding to the
+ − 2509
original definition. If there are variables, then we have to prefix this
+ − 2510
conversion with @{ML_ind fun_conv in Conv}. Now in Line 14 we only have to
+ − 2511
apply the new left-hand side to the generated conversion and obtain the the
+ − 2512
theorem we want to construct. As mentioned above, in Line 15 we only have to
565
+ − 2513
export the context \<open>ctxt'\<close> back to \<open>ctxt\<close> in order to
406
+ − 2514
produce meta-variables for the theorem. An example is as follows.
332
+ − 2515
+ − 2516
@{ML_response_fake [display, gray]
+ − 2517
"unabs_def @{context} @{thm id2_def}"
+ − 2518
"id2 ?x1 \<equiv> ?x1"}
565
+ − 2519
\<close>
+ − 2520
+ − 2521
text \<open>
243
+ − 2522
To sum up this section, conversions are more general than the simplifier
+ − 2523
or simprocs, but you have to do more work yourself. Also conversions are
+ − 2524
often much less efficient than the simplifier. The advantage of conversions,
406
+ − 2525
however, is that they provide much less room for non-termination.
146
+ − 2526
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2527
\begin{exercise}\label{ex:addconversion}
152
+ − 2528
Write a tactic that does the same as the simproc in exercise
291
+ − 2529
\ref{ex:addsimproc}, but is based on conversions. You can make
166
+ − 2530
the same assumptions as in \ref{ex:addsimproc}.
152
+ − 2531
\end{exercise}
+ − 2532
172
+ − 2533
\begin{exercise}\label{ex:compare}
174
+ − 2534
Compare your solutions of Exercises~\ref{ex:addsimproc} and \ref{ex:addconversion},
172
+ − 2535
and try to determine which way of rewriting such terms is faster. For this you might
+ − 2536
have to construct quite large terms. Also see Recipe \ref{rec:timing} for information
+ − 2537
about timing.
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2538
\end{exercise}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2539
146
+ − 2540
\begin{readmore}
424
+ − 2541
See @{ML_file "Pure/conv.ML"}
384
+ − 2542
for more information about conversion combinators.
243
+ − 2543
Some basic conversions are defined in @{ML_file "Pure/thm.ML"},
458
+ − 2544
@{ML_file "Pure/drule.ML"} and @{ML_file "Pure/raw_simplifier.ML"}.
146
+ − 2545
\end{readmore}
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2546
565
+ − 2547
\<close>
+ − 2548
+ − 2549
text \<open>
184
+ − 2550
(FIXME: check whether @{ML Pattern.match_rew} and @{ML Pattern.rewrite_term}
+ − 2551
are of any use/efficient)
565
+ − 2552
\<close>
+ − 2553
+ − 2554
+ − 2555
section \<open>Declarations (TBD)\<close>
+ − 2556
+ − 2557
section \<open>Structured Proofs\label{sec:structured} (TBD)\<close>
+ − 2558
+ − 2559
text \<open>TBD\<close>
129
+ − 2560
95
+ − 2561
lemma True
+ − 2562
proof
+ − 2563
+ − 2564
{
+ − 2565
fix A B C
+ − 2566
assume r: "A & B \<Longrightarrow> C"
+ − 2567
assume A B
+ − 2568
then have "A & B" ..
+ − 2569
then have C by (rule r)
+ − 2570
}
+ − 2571
+ − 2572
{
+ − 2573
fix A B C
+ − 2574
assume r: "A & B \<Longrightarrow> C"
+ − 2575
assume A B
+ − 2576
note conjI [OF this]
+ − 2577
note r [OF this]
+ − 2578
}
+ − 2579
oops
+ − 2580
565
+ − 2581
ML \<open>
95
+ − 2582
val ctxt0 = @{context};
+ − 2583
val ctxt = ctxt0;
+ − 2584
val (_, ctxt) = Variable.add_fixes ["A", "B", "C"] ctxt;
217
+ − 2585
val ([r], ctxt) = Assumption.add_assumes [@{cprop "A & B \<Longrightarrow> C"}] ctxt
+ − 2586
val (this, ctxt) = Assumption.add_assumes [@{cprop "A"}, @{cprop "B"}] ctxt;
95
+ − 2587
val this = [@{thm conjI} OF this];
+ − 2588
val this = r OF this;
+ − 2589
val this = Assumption.export false ctxt ctxt0 this
+ − 2590
val this = Variable.export ctxt ctxt0 [this]
565
+ − 2591
\<close>
+ − 2592
+ − 2593
section \<open>Summary\<close>
+ − 2594
+ − 2595
text \<open>
368
+ − 2596
+ − 2597
\begin{conventions}
+ − 2598
\begin{itemize}
565
+ − 2599
\item Reference theorems with the antiquotation \<open>@{thm \<dots>}\<close>.
368
+ − 2600
\item If a tactic is supposed to fail, return the empty sequence.
+ − 2601
\item If you apply a tactic to a sequence of subgoals, apply it
+ − 2602
in reverse order (i.e.~start with the last subgoal).
+ − 2603
\item Use simpsets that are as small as possible.
+ − 2604
\end{itemize}
+ − 2605
\end{conventions}
+ − 2606
565
+ − 2607
\<close>
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2608
139
+ − 2609
end