395
+ − 1
theory Advanced
441
+ − 2
imports Base First_Steps
318
+ − 3
begin
+ − 4
346
+ − 5
565
+ − 6
chapter \<open>Advanced Isabelle\label{chp:advanced}\<close>
381
+ − 7
565
+ − 8
text \<open>
410
+ − 9
\begin{flushright}
+ − 10
{\em All things are difficult before they are easy.} \\[1ex]
539
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 11
proverb\\[2ex]
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 12
{\em Programming is not just an act of telling a computer what
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 13
to do: it is also an act of telling other programmers what you
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 14
wished the computer to do. Both are important, and the latter
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 15
deserves care.}\\[1ex]
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 16
---Andrew Morton, Linux Kernel mailinglist, 13 March 2012
410
+ − 17
\end{flushright}
+ − 18
+ − 19
\medskip
400
+ − 20
While terms, types and theorems are the most basic data structures in
+ − 21
Isabelle, there are a number of layers built on top of them. Most of these
408
+ − 22
layers are concerned with storing and manipulating data. Handling them
487
+ − 23
properly is an essential skill for programming on the ML-level of Isabelle.
+ − 24
The most basic layer are theories. They contain global data and
408
+ − 25
can be seen as the ``long-term memory'' of Isabelle. There is usually only
+ − 26
one theory active at each moment. Proof contexts and local theories, on the
400
+ − 27
other hand, store local data for a task at hand. They act like the
408
+ − 28
``short-term memory'' and there can be many of them that are active in
+ − 29
parallel.
565
+ − 30
\<close>
318
+ − 31
565
+ − 32
section \<open>Theories\label{sec:theories}\<close>
358
+ − 33
565
+ − 34
text \<open>
492
+ − 35
Theories, as said above, are the most basic layer of abstraction in
+ − 36
Isabelle. They record information about definitions, syntax declarations, axioms,
488
+ − 37
theorems and much more. For example, if a definition is made, it
+ − 38
must be stored in a theory in order to be usable later on. Similar
+ − 39
with proofs: once a proof is finished, the proved theorem needs to
+ − 40
be stored in the theorem database of the theory in order to be
491
+ − 41
usable. All relevant data of a theory can be queried with the
+ − 42
Isabelle command \isacommand{print\_theory}.
358
+ − 43
+ − 44
\begin{isabelle}
+ − 45
\isacommand{print\_theory}\\
565
+ − 46
\<open>> names: Pure Code_Generator HOL \<dots>\<close>\\
+ − 47
\<open>> classes: Inf < type \<dots>\<close>\\
+ − 48
\<open>> default sort: type\<close>\\
+ − 49
\<open>> syntactic types: #prop \<dots>\<close>\\
+ − 50
\<open>> logical types: 'a \<times> 'b \<dots>\<close>\\
+ − 51
\<open>> type arities: * :: (random, random) random \<dots>\<close>\\
+ − 52
\<open>> logical constants: == :: 'a \<Rightarrow> 'a \<Rightarrow> prop \<dots>\<close>\\
+ − 53
\<open>> abbreviations: \<dots>\<close>\\
+ − 54
\<open>> axioms: \<dots>\<close>\\
+ − 55
\<open>> oracles: \<dots>\<close>\\
+ − 56
\<open>> definitions: \<dots>\<close>\\
+ − 57
\<open>> theorems: \<dots>\<close>
358
+ − 58
\end{isabelle}
+ − 59
565
+ − 60
Functions acting on theories often end with the suffix \<open>_global\<close>,
487
+ − 61
for example the function @{ML read_term_global in Syntax} in the structure
567
+ − 62
@{ML_structure Syntax}. The reason is to set them syntactically apart from
491
+ − 63
functions acting on contexts or local theories, which will be discussed in
+ − 64
the next sections. There is a tendency amongst Isabelle developers to prefer
487
+ − 65
``non-global'' operations, because they have some advantages, as we will also
490
+ − 66
discuss later. However, some basic understanding of theories is still necessary
+ − 67
for effective Isabelle programming.
487
+ − 68
491
+ − 69
An important Isabelle command with theories is \isacommand{setup}. In the
492
+ − 70
previous chapters we used it already to make a theorem attribute known
+ − 71
to Isabelle and to register a theorem under a name. What happens behind the
490
+ − 72
scenes is that \isacommand{setup} expects a function of type @{ML_type
+ − 73
"theory -> theory"}: the input theory is the current theory and the output
+ − 74
the theory where the attribute has been registered or the theorem has been
+ − 75
stored. This is a fundamental principle in Isabelle. A similar situation
491
+ − 76
arises with declaring a constant, which can be done on the ML-level with
567
+ − 77
function @{ML_ind declare_const in Sign} from the structure @{ML_structure
490
+ − 78
Sign}. To see how \isacommand{setup} works, consider the following code:
+ − 79
565
+ − 80
\<close>
348
+ − 81
565
+ − 82
ML %grayML\<open>let
485
+ − 83
val thy = @{theory}
+ − 84
val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
+ − 85
in
+ − 86
Sign.declare_const @{context} bar_const thy
565
+ − 87
end\<close>
348
+ − 88
565
+ − 89
text \<open>
488
+ − 90
If you simply run this code\footnote{Recall that ML-code needs to be enclosed in
565
+ − 91
\isacommand{ML}~\<open>\<verbopen> \<dots> \<verbclose>\<close>.} with the
+ − 92
intention of declaring a constant \<open>BAR\<close> having type @{typ nat}, then
488
+ − 93
indeed you obtain a theory as result. But if you query the
486
+ − 94
constant on the Isabelle level using the command \isacommand{term}
348
+ − 95
+ − 96
\begin{isabelle}
565
+ − 97
\isacommand{term}~\<open>BAR\<close>\\
+ − 98
\<open>> "BAR" :: "'a"\<close>
348
+ − 99
\end{isabelle}
+ − 100
495
+ − 101
you can see that you do \emph{not} obtain the expected constant of type @{typ nat}, but a free
484
+ − 102
variable (printed in blue) of polymorphic type. The problem is that the
+ − 103
ML-expression above did not ``register'' the declaration with the current theory.
+ − 104
This is what the command \isacommand{setup} is for. The constant is properly
+ − 105
declared with
565
+ − 106
\<close>
348
+ − 107
565
+ − 108
setup %gray \<open>fn thy =>
486
+ − 109
let
485
+ − 110
val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
486
+ − 111
val (_, thy') = Sign.declare_const @{context} bar_const thy
485
+ − 112
in
486
+ − 113
thy'
565
+ − 114
end\<close>
348
+ − 115
565
+ − 116
text \<open>
486
+ − 117
where the declaration is actually applied to the current theory and
348
+ − 118
+ − 119
\begin{isabelle}
+ − 120
\isacommand{term}~@{text [quotes] "BAR"}\\
565
+ − 121
\<open>> "BAR" :: "nat"\<close>
348
+ − 122
\end{isabelle}
+ − 123
492
+ − 124
now returns a (black) constant with the type @{typ nat}, as expected.
348
+ − 125
491
+ − 126
In a sense, \isacommand{setup} can be seen as a transaction that
565
+ − 127
takes the current theory \<open>thy\<close>, applies an operation, and
+ − 128
produces a new current theory \<open>thy'\<close>. This means that we have
491
+ − 129
to be careful to apply operations always to the most current theory,
+ − 130
not to a \emph{stale} one. Consider again the function inside the
486
+ − 131
\isacommand{setup}-command:
485
+ − 132
486
+ − 133
\begin{isabelle}
+ − 134
\begin{graybox}
569
+ − 135
\isacommand{setup}~\<open>\<verbopen>\<close> @{ML \<open>fn thy =>
486
+ − 136
let
569
+ − 137
val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
486
+ − 138
val (_, thy') = Sign.declare_const @{context} bar_const thy
+ − 139
in
+ − 140
thy
569
+ − 141
end\<close>}~\<open>\<verbclose>\<close>\isanewline
565
+ − 142
\<open>> ERROR: "Stale theory encountered"\<close>
486
+ − 143
\end{graybox}
+ − 144
\end{isabelle}
+ − 145
565
+ − 146
This time we erroneously return the original theory \<open>thy\<close>, instead of
+ − 147
the modified one \<open>thy'\<close>. Such buggy code will always result into
488
+ − 148
a runtime error message about stale theories.
486
+ − 149
502
+ − 150
\begin{readmore}
+ − 151
Most of the functions about theories are implemented in
+ − 152
@{ML_file "Pure/theory.ML"} and @{ML_file "Pure/global_theory.ML"}.
+ − 153
\end{readmore}
565
+ − 154
\<close>
318
+ − 155
565
+ − 156
section \<open>Contexts\<close>
341
+ − 157
565
+ − 158
text \<open>
486
+ − 159
Contexts are arguably more important than theories, even though they only
+ − 160
contain ``short-term memory data''. The reason is that a vast number of
490
+ − 161
functions in Isabelle depend in one way or another on contexts. Even such
+ − 162
mundane operations like printing out a term make essential use of contexts.
494
+ − 163
For this consider the following contrived proof-snippet whose purpose is to
490
+ − 164
fix two variables:
565
+ − 165
\<close>
486
+ − 166
488
+ − 167
lemma "True"
+ − 168
proof -
565
+ − 169
ML_prf \<open>Variable.dest_fixes @{context}\<close>
518
+ − 170
fix x y
565
+ − 171
ML_prf \<open>Variable.dest_fixes @{context}\<close>(*<*)oops(*>*)
488
+ − 172
565
+ − 173
text \<open>
496
+ − 174
The interesting point is that we injected ML-code before and after
490
+ − 175
the variables are fixed. For this remember that ML-code inside a proof
565
+ − 176
needs to be enclosed inside \isacommand{ML\_prf}~\<open>\<verbopen> \<dots> \<verbclose>\<close>,
+ − 177
not \isacommand{ML}~\<open>\<verbopen> \<dots> \<verbclose>\<close>. The function
567
+ − 178
@{ML_ind dest_fixes in Variable} from the structure @{ML_structure Variable} takes
490
+ − 179
a context and returns all its currently fixed variable (names). That
+ − 180
means a context has a dataslot containing information about fixed variables.
565
+ − 181
The ML-antiquotation \<open>@{context}\<close> points to the context that is
490
+ − 182
active at that point of the theory. Consequently, in the first call to
+ − 183
@{ML dest_fixes in Variable} this dataslot is empty; in the second it is
565
+ − 184
filled with \<open>x\<close> and \<open>y\<close>. What is interesting is that contexts
494
+ − 185
can be stacked. For this consider the following proof fragment:
565
+ − 186
\<close>
490
+ − 187
+ − 188
lemma "True"
+ − 189
proof -
+ − 190
fix x y
+ − 191
{ fix z w
565
+ − 192
ML_prf \<open>Variable.dest_fixes @{context}\<close>
518
+ − 193
}
565
+ − 194
ML_prf \<open>Variable.dest_fixes @{context}\<close>(*<*)oops(*>*)
491
+ − 195
565
+ − 196
text \<open>
495
+ − 197
Here the first time we call @{ML dest_fixes in Variable} we have four fixes variables;
565
+ − 198
the second time we get only the fixes variables \<open>x\<close> and \<open>y\<close> as answer, since
+ − 199
\<open>z\<close> and \<open>w\<close> are not anymore in the scope of the context.
495
+ − 200
This means the curly-braces act on the Isabelle level as opening and closing statements
496
+ − 201
for a context. The above proof fragment corresponds roughly to the following ML-code
565
+ − 202
\<close>
488
+ − 203
565
+ − 204
ML %grayML\<open>val ctxt0 = @{context};
492
+ − 205
val ([x, y], ctxt1) = Variable.add_fixes ["x", "y"] ctxt0;
565
+ − 206
val ([z, w], ctxt2) = Variable.add_fixes ["z", "w"] ctxt1\<close>
492
+ − 207
565
+ − 208
text \<open>
494
+ − 209
where the function @{ML_ind add_fixes in Variable} fixes a list of variables
+ − 210
specified by strings. Let us come back to the point about printing terms. Consider
565
+ − 211
printing out the term \mbox{\<open>(x, y, z, w)\<close>} using our function @{ML_ind pretty_term}.
496
+ − 212
This function takes a term and a context as argument. Notice how the printing
498
+ − 213
of the term changes according to which context is used.
492
+ − 214
495
+ − 215
\begin{isabelle}
+ − 216
\begin{graybox}
569
+ − 217
@{ML \<open>let
+ − 218
val trm = @{term "(x, y, z, w)"}
493
+ − 219
in
+ − 220
pwriteln (Pretty.chunks
+ − 221
[ pretty_term ctxt0 trm,
+ − 222
pretty_term ctxt1 trm,
+ − 223
pretty_term ctxt2 trm ])
569
+ − 224
end\<close>}\\
495
+ − 225
\setlength{\fboxsep}{0mm}
565
+ − 226
\<open>>\<close>~\<open>(\<close>\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>x\<close>}}\<open>,\<close>~%
+ − 227
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>y\<close>}}\<open>,\<close>~%
+ − 228
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>z\<close>}}\<open>,\<close>~%
+ − 229
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>w\<close>}}\<open>)\<close>\\
+ − 230
\<open>>\<close>~\<open>(\<close>\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>x\<close>}}\<open>,\<close>~%
+ − 231
\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>y\<close>}}\<open>,\<close>~%
+ − 232
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>z\<close>}}\<open>,\<close>~%
+ − 233
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>w\<close>}}\<open>)\<close>\\
+ − 234
\<open>>\<close>~\<open>(\<close>\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>x\<close>}}\<open>,\<close>~%
+ − 235
\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>y\<close>}}\<open>,\<close>~%
+ − 236
\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>z\<close>}}\<open>,\<close>~%
+ − 237
\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>w\<close>}}\<open>)\<close>
495
+ − 238
\end{graybox}
+ − 239
\end{isabelle}
+ − 240
+ − 241
+ − 242
The term we are printing is in all three cases the same---a tuple containing
569
+ − 243
four free variables. As you can see, however, in case of @{ML \<open>ctxt0\<close>} all
+ − 244
variables are highlighted indicating a problem, while in case of @{ML \<open>ctxt1\<close>} \<open>x\<close> and \<open>y\<close> are printed as normal (blue) free
565
+ − 245
variables, but not \<open>z\<close> and \<open>w\<close>. In the last case all variables
495
+ − 246
are printed as expected. The point of this example is that the context
+ − 247
contains the information which variables are fixed, and designates all other
497
+ − 248
free variables as being alien or faulty. Therefore the highlighting.
+ − 249
While this seems like a minor detail, the concept of making the context aware
+ − 250
of fixed variables is actually quite useful. For example it prevents us from
501
+ − 251
fixing a variable twice
495
+ − 252
572
+ − 253
@{ML_response [gray, display]
569
+ − 254
\<open>@{context}
+ − 255
|> Variable.add_fixes ["x", "x"]\<close>
572
+ − 256
\<open>Duplicate fixed variable(s): "x"\<close>}
495
+ − 257
501
+ − 258
More importantly it also allows us to easily create fresh names for
+ − 259
fixed variables. For this you have to use the function @{ML_ind
567
+ − 260
variant_fixes in Variable} from the structure @{ML_structure Variable}.
501
+ − 261
572
+ − 262
@{ML_response [gray, display]
569
+ − 263
\<open>@{context}
+ − 264
|> Variable.variant_fixes ["y", "y", "z"]\<close>
572
+ − 265
\<open>(["y", "ya", "z"],\<dots>\<close>}
501
+ − 266
565
+ − 267
Now a fresh variant for the second occurence of \<open>y\<close> is created
502
+ − 268
avoiding any clash. In this way we can also create fresh free variables
501
+ − 269
that avoid any clashes with fixed variables. In Line~3 below we fix
565
+ − 270
the variable \<open>x\<close> in the context \<open>ctxt1\<close>. Next we want to
501
+ − 271
create two fresh variables of type @{typ nat} as variants of the
+ − 272
string @{text [quotes] "x"} (Lines 6 and 7).
495
+ − 273
572
+ − 274
@{ML_matchresult [display, gray, linenos]
569
+ − 275
\<open>let
495
+ − 276
val ctxt0 = @{context}
569
+ − 277
val (_, ctxt1) = Variable.add_fixes ["x"] ctxt0
+ − 278
val frees = replicate 2 ("x", @{typ nat})
495
+ − 279
in
+ − 280
(Variable.variant_frees ctxt0 [] frees,
+ − 281
Variable.variant_frees ctxt1 [] frees)
569
+ − 282
end\<close>
572
+ − 283
\<open>([("x", _), ("xa", _)],
+ − 284
[("xa", _), ("xb", _)])\<close>}
495
+ − 285
565
+ − 286
As you can see, if we create the fresh variables with the context \<open>ctxt0\<close>,
+ − 287
then we obtain \<open>x\<close> and \<open>xa\<close>; but in \<open>ctxt1\<close> we obtain \<open>xa\<close>
+ − 288
and \<open>xb\<close> avoiding \<open>x\<close>, which was fixed in this context.
495
+ − 289
496
+ − 290
Often one has the problem that given some terms, create fresh variables
+ − 291
avoiding any variable occurring in those terms. For this you can use the
567
+ − 292
function @{ML_ind declare_term in Variable} from the structure @{ML_structure Variable}.
495
+ − 293
572
+ − 294
@{ML_matchresult [gray, display]
569
+ − 295
\<open>let
+ − 296
val ctxt1 = Variable.declare_term @{term "(x, xa)"} @{context}
+ − 297
val frees = replicate 2 ("x", @{typ nat})
495
+ − 298
in
+ − 299
Variable.variant_frees ctxt1 [] frees
569
+ − 300
end\<close>
572
+ − 301
\<open>[("xb", _), ("xc", _)]\<close>}
495
+ − 302
565
+ − 303
The result is \<open>xb\<close> and \<open>xc\<close> for the names of the fresh
+ − 304
variables, since \<open>x\<close> and \<open>xa\<close> occur in the term we declared.
498
+ − 305
Note that @{ML_ind declare_term in Variable} does not fix the
+ − 306
variables; it just makes them ``known'' to the context. You can see
+ − 307
that if you print out a declared term.
+ − 308
+ − 309
\begin{isabelle}
+ − 310
\begin{graybox}
569
+ − 311
@{ML \<open>let
+ − 312
val trm = @{term "P x y z"}
498
+ − 313
val ctxt1 = Variable.declare_term trm @{context}
+ − 314
in
+ − 315
pwriteln (pretty_term ctxt1 trm)
569
+ − 316
end\<close>}\\
498
+ − 317
\setlength{\fboxsep}{0mm}
565
+ − 318
\<open>>\<close>~\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>P\<close>}}~%
+ − 319
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>x\<close>}}~%
+ − 320
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>y\<close>}}~%
+ − 321
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>z\<close>}}
498
+ − 322
\end{graybox}
+ − 323
\end{isabelle}
+ − 324
+ − 325
All variables are highligted, indicating that they are not
+ − 326
fixed. However, declaring a term is helpful when parsing terms using
+ − 327
the function @{ML_ind read_term in Syntax} from the structure
567
+ − 328
@{ML_structure Syntax}. Consider the following code:
495
+ − 329
572
+ − 330
@{ML_response [gray, display]
569
+ − 331
\<open>let
495
+ − 332
val ctxt0 = @{context}
569
+ − 333
val ctxt1 = Variable.declare_term @{term "x::nat"} ctxt0
495
+ − 334
in
569
+ − 335
(Syntax.read_term ctxt0 "x",
+ − 336
Syntax.read_term ctxt1 "x")
+ − 337
end\<close>
+ − 338
\<open>(Free ("x", "'a"), Free ("x", "nat"))\<close>}
495
+ − 339
565
+ − 340
Parsing the string in the context \<open>ctxt0\<close> results in a free variable
+ − 341
with a default polymorphic type, but in case of \<open>ctxt1\<close> we obtain a
495
+ − 342
free variable of type @{typ nat} as previously declared. Which
496
+ − 343
type the parsed term receives depends upon the last declaration that
+ − 344
is made, as the next example illustrates.
495
+ − 345
572
+ − 346
@{ML_response [gray, display]
569
+ − 347
\<open>let
+ − 348
val ctxt1 = Variable.declare_term @{term "x::nat"} @{context}
+ − 349
val ctxt2 = Variable.declare_term @{term "x::int"} ctxt1
495
+ − 350
in
569
+ − 351
(Syntax.read_term ctxt1 "x",
+ − 352
Syntax.read_term ctxt2 "x")
+ − 353
end\<close>
+ − 354
\<open>(Free ("x", "nat"), Free ("x", "int"))\<close>}
495
+ − 355
499
+ − 356
The most useful feature of contexts is that one can export, or transfer,
+ − 357
terms and theorems between them. We show this first for terms.
497
+ − 358
+ − 359
\begin{isabelle}
+ − 360
\begin{graybox}
+ − 361
\begin{linenos}
569
+ − 362
@{ML \<open>let
497
+ − 363
val ctxt0 = @{context}
569
+ − 364
val (_, ctxt1) = Variable.add_fixes ["x", "y", "z"] ctxt0
+ − 365
val foo_trm = @{term "P x y z"}
497
+ − 366
in
+ − 367
singleton (Variable.export_terms ctxt1 ctxt0) foo_trm
+ − 368
|> pretty_term ctxt0
+ − 369
|> pwriteln
569
+ − 370
end\<close>}
497
+ − 371
\end{linenos}
+ − 372
\setlength{\fboxsep}{0mm}
565
+ − 373
\<open>>\<close>~\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>P\<close>}}~%
+ − 374
\<open>?x ?y ?z\<close>
497
+ − 375
\end{graybox}
+ − 376
\end{isabelle}
+ − 377
498
+ − 378
In Line 3 we fix the variables @{term x}, @{term y} and @{term z} in
565
+ − 379
context \<open>ctxt1\<close>. The function @{ML_ind export_terms in
567
+ − 380
Variable} from the structure @{ML_structure Variable} can be used to transfer
498
+ − 381
terms between contexts. Transferring means to turn all (free)
+ − 382
variables that are fixed in one context, but not in the other, into
+ − 383
schematic variables. In our example, we are transferring the term
565
+ − 384
\<open>P x y z\<close> from context \<open>ctxt1\<close> to \<open>ctxt0\<close>,
498
+ − 385
which means @{term x}, @{term y} and @{term z} become schematic
500
+ − 386
variables (as can be seen by the leading question marks in the result).
565
+ − 387
Note that the variable \<open>P\<close> stays a free variable, since it not fixed in
+ − 388
\<open>ctxt1\<close>; it is even highlighed, because \<open>ctxt0\<close> does
498
+ − 389
not know about it. Note also that in Line 6 we had to use the
+ − 390
function @{ML_ind singleton}, because the function @{ML_ind
+ − 391
export_terms in Variable} normally works over lists of terms.
+ − 392
+ − 393
The case of transferring theorems is even more useful. The reason is
+ − 394
that the generalisation of fixed variables to schematic variables is
499
+ − 395
not trivial if done manually. For illustration purposes we use in the
+ − 396
following code the function @{ML_ind make_thm in Skip_Proof} from the
567
+ − 397
structure @{ML_structure Skip_Proof}. This function will turn an arbitray
500
+ − 398
term, in our case @{term "P x y z x y z"}, into a theorem (disregarding
+ − 399
whether it is actually provable).
498
+ − 400
572
+ − 401
@{ML_response [display, gray]
569
+ − 402
\<open>let
498
+ − 403
val thy = @{theory}
+ − 404
val ctxt0 = @{context}
569
+ − 405
val (_, ctxt1) = Variable.add_fixes ["P", "x", "y", "z"] ctxt0
+ − 406
val foo_thm = Skip_Proof.make_thm thy @{prop "P x y z x y z"}
498
+ − 407
in
+ − 408
singleton (Proof_Context.export ctxt1 ctxt0) foo_thm
569
+ − 409
end\<close>
+ − 410
\<open>?P ?x ?y ?z ?x ?y ?z\<close>}
499
+ − 411
565
+ − 412
Since we fixed all variables in \<open>ctxt1\<close>, in the exported
502
+ − 413
result all of them are schematic. The great point of contexts is
+ − 414
that exporting from one to another is not just restricted to
+ − 415
variables, but also works with assumptions. For this we can use the
+ − 416
function @{ML_ind export in Assumption} from the structure
567
+ − 417
@{ML_structure Assumption}. Consider the following code.
500
+ − 418
572
+ − 419
@{ML_response [display, gray, linenos]
569
+ − 420
\<open>let
500
+ − 421
val ctxt0 = @{context}
569
+ − 422
val ([eq], ctxt1) = Assumption.add_assumes [@{cprop "x \<equiv> y"}] ctxt0
500
+ − 423
val eq' = Thm.symmetric eq
+ − 424
in
+ − 425
Assumption.export false ctxt1 ctxt0 eq'
569
+ − 426
end\<close>
+ − 427
\<open>x \<equiv> y \<Longrightarrow> y \<equiv> x\<close>}
500
+ − 428
+ − 429
The function @{ML_ind add_assumes in Assumption} from the structure
567
+ − 430
@{ML_structure Assumption} adds the assumption \mbox{\<open>x \<equiv> y\<close>}
565
+ − 431
to the context \<open>ctxt1\<close> (Line 3). This function expects a list
500
+ − 432
of @{ML_type cterm}s and returns them as theorems, together with the
+ − 433
new context in which they are ``assumed''. In Line 4 we use the
567
+ − 434
function @{ML_ind symmetric in Thm} from the structure @{ML_structure
500
+ − 435
Thm} in order to obtain the symmetric version of the assumed
565
+ − 436
meta-equality. Now exporting the theorem \<open>eq'\<close> from \<open>ctxt1\<close> to \<open>ctxt0\<close> means @{term "y \<equiv> x"} will be prefixed with
500
+ − 437
the assumed theorem. The boolean flag in @{ML_ind export in
+ − 438
Assumption} indicates whether the assumptions should be marked with
+ − 439
the goal marker (see Section~\ref{sec:basictactics}). In normal
+ − 440
circumstances this is not necessary and so should be set to @{ML
+ − 441
false}. The result of the export is then the theorem \mbox{@{term
+ − 442
"x \<equiv> y \<Longrightarrow> y \<equiv> x"}}. As can be seen this is an easy way for obtaing
+ − 443
simple theorems. We will explain this in more detail in
+ − 444
Section~\ref{sec:structured}.
+ − 445
+ − 446
The function @{ML_ind export in Proof_Context} from the structure
567
+ − 447
@{ML_structure Proof_Context} combines both export functions from
+ − 448
@{ML_structure Variable} and @{ML_structure Assumption}. This can be seen
500
+ − 449
in the following example.
+ − 450
572
+ − 451
@{ML_response [display, gray]
569
+ − 452
\<open>let
500
+ − 453
val ctxt0 = @{context}
+ − 454
val ((fvs, [eq]), ctxt1) = ctxt0
569
+ − 455
|> Variable.add_fixes ["x", "y"]
+ − 456
||>> Assumption.add_assumes [@{cprop "x \<equiv> y"}]
500
+ − 457
val eq' = Thm.symmetric eq
+ − 458
in
+ − 459
Proof_Context.export ctxt1 ctxt0 [eq']
569
+ − 460
end\<close>
572
+ − 461
\<open>["?x \<equiv> ?y \<Longrightarrow> ?y \<equiv> ?x"]\<close>}
565
+ − 462
\<close>
495
+ − 463
496
+ − 464
493
+ − 465
565
+ − 466
text \<open>
493
+ − 467
565
+ − 468
\<close>
493
+ − 469
492
+ − 470
486
+ − 471
(*
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 472
ML %grayML{*Proof_Context.debug := true*}
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 473
ML %grayML{*Proof_Context.verbose := true*}
486
+ − 474
*)
+ − 475
487
+ − 476
(*
486
+ − 477
lemma "True"
+ − 478
proof -
+ − 479
{ -- "\<And>x. _"
+ − 480
fix x
+ − 481
have "B x" sorry
+ − 482
thm this
+ − 483
}
+ − 484
+ − 485
thm this
+ − 486
+ − 487
{ -- "A \<Longrightarrow> _"
+ − 488
assume A
+ − 489
have B sorry
+ − 490
thm this
+ − 491
}
+ − 492
+ − 493
thm this
+ − 494
+ − 495
{ -- "\<And>x. x = _ \<Longrightarrow> _"
+ − 496
def x \<equiv> a
+ − 497
have "B x" sorry
+ − 498
}
+ − 499
+ − 500
thm this
+ − 501
+ − 502
oops
487
+ − 503
*)
413
+ − 504
565
+ − 505
section \<open>Local Theories and Local Setups\label{sec:local} (TBD)\<close>
341
+ − 506
565
+ − 507
text \<open>
400
+ − 508
In contrast to an ordinary theory, which simply consists of a type
+ − 509
signature, as well as tables for constants, axioms and theorems, a local
+ − 510
theory contains additional context information, such as locally fixed
+ − 511
variables and local assumptions that may be used by the package. The type
+ − 512
@{ML_type local_theory} is identical to the type of \emph{proof contexts}
+ − 513
@{ML_type "Proof.context"}, although not every proof context constitutes a
+ − 514
valid local theory.
+ − 515
569
+ − 516
@{ML \<open>Context.>> o Context.map_theory\<close>}
394
+ − 517
@{ML_ind "Local_Theory.declaration"}
486
+ − 518
+ − 519
A similar command is \isacommand{local\_setup}, which expects a function
+ − 520
of type @{ML_type "local_theory -> local_theory"}. Later on we will also
+ − 521
use the commands \isacommand{method\_setup} for installing methods in the
+ − 522
current theory and \isacommand{simproc\_setup} for adding new simprocs to
+ − 523
the current simpset.
565
+ − 524
\<close>
318
+ − 525
336
+ − 526
565
+ − 527
section \<open>Morphisms (TBD)\<close>
394
+ − 528
565
+ − 529
text \<open>
394
+ − 530
Morphisms are arbitrary transformations over terms, types, theorems and bindings.
+ − 531
They can be constructed using the function @{ML_ind morphism in Morphism},
+ − 532
which expects a record with functions of type
+ − 533
+ − 534
\begin{isabelle}
+ − 535
\begin{tabular}{rl}
565
+ − 536
\<open>binding:\<close> & \<open>binding -> binding\<close>\\
+ − 537
\<open>typ:\<close> & \<open>typ -> typ\<close>\\
+ − 538
\<open>term:\<close> & \<open>term -> term\<close>\\
+ − 539
\<open>fact:\<close> & \<open>thm list -> thm list\<close>
394
+ − 540
\end{tabular}
+ − 541
\end{isabelle}
+ − 542
+ − 543
The simplest morphism is the @{ML_ind identity in Morphism}-morphism defined as
565
+ − 544
\<close>
394
+ − 545
565
+ − 546
ML %grayML\<open>val identity = Morphism.morphism "" {binding = [], typ = [], term = [], fact = []}\<close>
394
+ − 547
565
+ − 548
text \<open>
394
+ − 549
Morphisms can be composed with the function @{ML_ind "$>" in Morphism}
565
+ − 550
\<close>
394
+ − 551
565
+ − 552
ML %grayML\<open>fun trm_phi (Free (x, T)) = Var ((x, 0), T)
394
+ − 553
| trm_phi (Abs (x, T, t)) = Abs (x, T, trm_phi t)
+ − 554
| trm_phi (t $ s) = (trm_phi t) $ (trm_phi s)
565
+ − 555
| trm_phi t = t\<close>
394
+ − 556
565
+ − 557
ML %grayML\<open>val phi = Morphism.term_morphism "foo" trm_phi\<close>
394
+ − 558
565
+ − 559
ML %grayML\<open>Morphism.term phi @{term "P x y"}\<close>
394
+ − 560
565
+ − 561
text \<open>
394
+ − 562
@{ML_ind term_morphism in Morphism}
+ − 563
+ − 564
@{ML_ind term in Morphism},
+ − 565
@{ML_ind thm in Morphism}
+ − 566
+ − 567
\begin{readmore}
+ − 568
Morphisms are implemented in the file @{ML_file "Pure/morphism.ML"}.
+ − 569
\end{readmore}
565
+ − 570
\<close>
318
+ − 571
565
+ − 572
section \<open>Misc (TBD)\<close>
318
+ − 573
565
+ − 574
text \<open>
319
+ − 575
FIXME: association lists:
+ − 576
@{ML_file "Pure/General/alist.ML"}
+ − 577
+ − 578
FIXME: calling the ML-compiler
+ − 579
565
+ − 580
\<close>
319
+ − 581
565
+ − 582
section \<open>What Is In an Isabelle Name? (TBD)\<close>
414
+ − 583
565
+ − 584
text \<open>
414
+ − 585
On the ML-level of Isabelle, you often have to work with qualified names.
+ − 586
These are strings with some additional information, such as positional
+ − 587
information and qualifiers. Such qualified names can be generated with the
565
+ − 588
antiquotation \<open>@{binding \<dots>}\<close>. For example
414
+ − 589
567
+ − 590
@{ML_matchresult [display,gray]
569
+ − 591
\<open>@{binding "name"}\<close>
+ − 592
\<open>name\<close>}
414
+ − 593
+ − 594
An example where a qualified name is needed is the function
+ − 595
@{ML_ind define in Local_Theory}. This function is used below to define
+ − 596
the constant @{term "TrueConj"} as the conjunction @{term "True \<and> True"}.
565
+ − 597
\<close>
414
+ − 598
565
+ − 599
local_setup %gray \<open>
414
+ − 600
Local_Theory.define ((@{binding "TrueConj"}, NoSyn),
565
+ − 601
((@{binding "TrueConj_def"}, []), @{term "True \<and> True"})) #> snd\<close>
414
+ − 602
565
+ − 603
text \<open>
414
+ − 604
Now querying the definition you obtain:
+ − 605
+ − 606
\begin{isabelle}
565
+ − 607
\isacommand{thm}~\<open>TrueConj_def\<close>\\
+ − 608
\<open>> \<close>~@{thm TrueConj_def}
414
+ − 609
\end{isabelle}
+ − 610
+ − 611
\begin{readmore}
+ − 612
The basic operations on bindings are implemented in
+ − 613
@{ML_file "Pure/General/binding.ML"}.
+ − 614
\end{readmore}
+ − 615
+ − 616
\footnote{\bf FIXME give a better example why bindings are important}
+ − 617
\footnote{\bf FIXME give a pointer to \isacommand{local\_setup}; if not, then explain
+ − 618
why @{ML snd} is needed.}
+ − 619
\footnote{\bf FIXME: There should probably a separate section on binding, long-names
+ − 620
and sign.}
+ − 621
565
+ − 622
\<close>
414
+ − 623
318
+ − 624
565
+ − 625
ML \<open>Sign.intern_type @{theory} "list"\<close>
+ − 626
ML \<open>Sign.intern_const @{theory} "prod_fun"\<close>
360
+ − 627
565
+ − 628
text \<open>
414
+ − 629
\footnote{\bf FIXME: Explain the following better; maybe put in a separate
+ − 630
section and link with the comment in the antiquotation section.}
+ − 631
+ − 632
Occasionally you have to calculate what the ``base'' name of a given
462
+ − 633
constant is. For this you can use the function @{ML_ind Long_Name.base_name}. For example:
414
+ − 634
569
+ − 635
@{ML_matchresult [display,gray] \<open>Long_Name.base_name "List.list.Nil"\<close> \<open>"Nil"\<close>}
414
+ − 636
+ − 637
\begin{readmore}
+ − 638
Functions about naming are implemented in @{ML_file "Pure/General/name_space.ML"};
+ − 639
functions about signatures in @{ML_file "Pure/sign.ML"}.
+ − 640
\end{readmore}
565
+ − 641
\<close>
387
+ − 642
565
+ − 643
text \<open>
387
+ − 644
@{ML_ind "Binding.name_of"} returns the string without markup
394
+ − 645
562
+ − 646
@{ML_ind "Binding.concealed"}
565
+ − 647
\<close>
387
+ − 648
565
+ − 649
section \<open>Concurrency (TBD)\<close>
388
+ − 650
565
+ − 651
text \<open>
388
+ − 652
@{ML_ind prove_future in Goal}
+ − 653
@{ML_ind future_result in Goal}
565
+ − 654
\<close>
387
+ − 655
565
+ − 656
section \<open>Parse and Print Translations (TBD)\<close>
396
+ − 657
565
+ − 658
section \<open>Summary\<close>
349
+ − 659
565
+ − 660
text \<open>
395
+ − 661
TBD
565
+ − 662
\<close>
318
+ − 663
+ − 664
end