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