441
+ − 1
theory First_Steps
25
+ − 2
imports Base
2
+ − 3
begin
557
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 4
565
+ − 5
chapter \<open>First Steps\label{chp:firststeps}\<close>
2
+ − 6
565
+ − 7
text \<open>
489
+ − 8
\begin{flushright}
+ − 9
{\em ``The most effective debugging tool is still careful thought,\\
+ − 10
coupled with judiciously placed print statements.''} \\[1ex]
+ − 11
Brian Kernighan, in {\em Unix for Beginners}, 1979
323
+ − 12
\end{flushright}
+ − 13
324
+ − 14
\medskip
562
+ − 15
Isabelle programming is done in ML @{cite "isa-imp"}. Just like lemmas and proofs, ML-code for
311
+ − 16
Isabelle must be part of a theory. If you want to follow the code given in
54
+ − 17
this chapter, we assume you are working inside the theory starting with
2
+ − 18
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 19
\begin{quote}
5
+ − 20
\begin{tabular}{@ {}l}
441
+ − 21
\isacommand{theory} First\_Steps\\
5
+ − 22
\isacommand{imports} Main\\
+ − 23
\isacommand{begin}\\
6
+ − 24
\ldots
5
+ − 25
\end{tabular}
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 26
\end{quote}
157
+ − 27
324
+ − 28
We also generally assume you are working with the logic HOL. The examples
+ − 29
that will be given might need to be adapted if you work in a different logic.
565
+ − 30
\<close>
5
+ − 31
565
+ − 32
section \<open>Including ML-Code\label{sec:include}\<close>
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 33
565
+ − 34
text \<open>
311
+ − 35
The easiest and quickest way to include code in a theory is by using the
+ − 36
\isacommand{ML}-command. For example:
+ − 37
+ − 38
\begin{isabelle}
+ − 39
\begin{graybox}
565
+ − 40
\isacommand{ML}~\<open>\<verbopen>\<close>\isanewline
569
+ − 41
\hspace{5mm}@{ML \<open>3 + 4\<close>}\isanewline
565
+ − 42
\<open>\<verbclose>\<close>\isanewline
+ − 43
\<open>> 7\<close>\smallskip
311
+ − 44
\end{graybox}
+ − 45
\end{isabelle}
+ − 46
567
+ − 47
If you work with the
+ − 48
Isabelle/jEdit, then you just have to hover the cursor over the code
474
+ − 49
and you see the evaluated result in the ``Output'' window.
193
+ − 50
565
+ − 51
As mentioned in the Introduction, we will drop the \isacommand{ML}~\<open>\<verbopen> \<dots> \<verbclose>\<close> scaffolding whenever we show code. The lines
311
+ − 52
prefixed with @{text [quotes] ">"} are not part of the code, rather they
+ − 53
indicate what the response is when the code is evaluated. There are also
+ − 54
the commands \isacommand{ML\_val} and \isacommand{ML\_prf} for including
+ − 55
ML-code. The first evaluates the given code, but any effect on the theory,
+ − 56
in which the code is embedded, is suppressed. The second needs to be used if
+ − 57
ML-code is defined inside a proof. For example
253
+ − 58
254
+ − 59
\begin{quote}
+ − 60
\begin{isabelle}
565
+ − 61
\isacommand{lemma}~\<open>test:\<close>\isanewline
254
+ − 62
\isacommand{shows}~@{text [quotes] "True"}\isanewline
569
+ − 63
\isacommand{ML\_prf}~\<open>\<verbopen>\<close>~@{ML \<open>writeln "Trivial!"\<close>}~\<open>\<verbclose>\<close>\isanewline
254
+ − 64
\isacommand{oops}
+ − 65
\end{isabelle}
+ − 66
\end{quote}
253
+ − 67
502
+ − 68
However, both commands will only play minor roles in this tutorial (we most of
+ − 69
the time make sure that the ML-code is defined outside proofs).
10
+ − 70
538
+ − 71
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
+ − 72
Once a portion of code is relatively stable, you usually want to export it
235
+ − 73
to a separate ML-file. Such files can then be included somewhere inside a
538
+ − 74
theory by using the command \isacommand{ML\_file}. For example
235
+ − 75
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 76
\begin{quote}
235
+ − 77
\begin{tabular}{@ {}l}
441
+ − 78
\isacommand{theory} First\_Steps\\
235
+ − 79
\isacommand{imports} Main\\
+ − 80
\isacommand{begin}\\
+ − 81
\ldots\\
565
+ − 82
\isacommand{ML\_file}~\<open>"file_to_be_included.ML"\<close>\\
235
+ − 83
\ldots
+ − 84
\end{tabular}
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 85
\end{quote}
565
+ − 86
\<close>
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 87
565
+ − 88
section \<open>Printing and Debugging\label{sec:printing}\<close>
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 89
565
+ − 90
text \<open>
374
+ − 91
During development you might find it necessary to inspect data in your
311
+ − 92
code. This can be done in a ``quick-and-dirty'' fashion using the function
369
+ − 93
@{ML_ind writeln in Output}. For example
10
+ − 94
569
+ − 95
@{ML_matchresult_fake [display,gray] \<open>writeln "any string"\<close> \<open>"any string"\<close>}
10
+ − 96
557
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 97
will print out @{text [quotes] "any string"}.
474
+ − 98
This function expects a string as argument. If you develop under
311
+ − 99
PolyML, then there is a convenient, though again ``quick-and-dirty'', method
421
+ − 100
for converting values into strings, namely the antiquotation
565
+ − 101
\<open>@{make_string}\<close>:
10
+ − 102
567
+ − 103
@{ML_matchresult_fake [display,gray] \<open>writeln (@{make_string} 1)\<close> \<open>"1"\<close>}
12
+ − 104
565
+ − 105
However, \<open>@{make_string}\<close> only works if the type of what
466
+ − 106
is converted is monomorphic and not a function.
311
+ − 107
466
+ − 108
You can print out error messages with the function @{ML_ind error in Library};
+ − 109
for example:
75
+ − 110
572
+ − 111
@{ML_response [display,gray]
569
+ − 112
\<open>if 0 = 1 then true else (error "foo")\<close>
572
+ − 113
\<open>foo\<close>}
75
+ − 114
565
+ − 115
This function raises the exception \<open>ERROR\<close>, which will then
557
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 116
be displayed by the infrastructure indicating that it is an error by
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 117
painting the output red. Note that this exception is meant
466
+ − 118
for ``user-level'' error messages seen by the ``end-user''.
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 119
For messages where you want to indicate a genuine program error
565
+ − 120
use the exception \<open>Fail\<close>.
304
+ − 121
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 122
Most often you want to inspect contents of Isabelle's basic data structures,
414
+ − 123
namely @{ML_type term}, @{ML_type typ}, @{ML_type cterm}, @{ML_type ctyp}
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 124
and @{ML_type thm}. Isabelle provides elaborate pretty-printing functions,
440
+ − 125
which we will explain in more detail in Section \ref{sec:pretty}. For now
+ − 126
we just use the functions @{ML_ind writeln in Pretty} from the structure
567
+ − 127
@{ML_structure Pretty} and @{ML_ind pretty_term in Syntax} from the structure
+ − 128
@{ML_structure Syntax}. For more convenience, we bind them to the toplevel.
565
+ − 129
\<close>
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 130
565
+ − 131
ML %grayML\<open>val pretty_term = Syntax.pretty_term
+ − 132
val pwriteln = Pretty.writeln\<close>
572
+ − 133
(* We redfine pwriteln to return a value not just a side effect on the output in order to
+ − 134
use some checking of output with ML_response antiquotation. *)
+ − 135
ML %invisible\<open>val pretty_term = Syntax.pretty_term
+ − 136
val pwriteln = YXML.content_of o Pretty.string_of\<close>
565
+ − 137
text \<open>
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 138
They can be used as follows
126
+ − 139
572
+ − 140
@{ML_response [display,gray]
569
+ − 141
\<open>pwriteln (pretty_term @{context} @{term "1::nat"})\<close>
+ − 142
\<open>"1"\<close>}
126
+ − 143
440
+ − 144
If there is more than one term to be printed, you can use the
446
+ − 145
function @{ML_ind commas in Pretty} and @{ML_ind block in Pretty}
+ − 146
to separate them.
565
+ − 147
\<close>
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 148
565
+ − 149
ML %grayML\<open>fun pretty_terms ctxt trms =
+ − 150
Pretty.block (Pretty.commas (map (pretty_term ctxt) trms))\<close>
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 151
565
+ − 152
text \<open>
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 153
To print out terms together with their typing information,
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 154
set the configuration value
451
+ − 155
@{ML_ind show_types in Syntax} to @{ML true}.
565
+ − 156
\<close>
370
+ − 157
565
+ − 158
ML %grayML\<open>val show_types_ctxt = Config.put show_types true @{context}\<close>
370
+ − 159
565
+ − 160
text \<open>
451
+ − 161
Now by using this context @{ML pretty_term} prints out
370
+ − 162
572
+ − 163
@{ML_response [display, gray]
569
+ − 164
\<open>pwriteln (pretty_term show_types_ctxt @{term "(1::nat, x)"})\<close>
+ − 165
\<open>(1::nat, x::'a)\<close>}
370
+ − 166
565
+ − 167
where \<open>1\<close> and \<open>x\<close> are displayed with their inferred types.
546
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 168
Other configuration values that influence
502
+ − 169
printing of terms include
451
+ − 170
502
+ − 171
\begin{itemize}
+ − 172
\item @{ML_ind show_brackets in Syntax}
+ − 173
\item @{ML_ind show_sorts in Syntax}
+ − 174
\item @{ML_ind eta_contract in Syntax}
+ − 175
\end{itemize}
+ − 176
440
+ − 177
A @{ML_type cterm} can be printed with the following function.
565
+ − 178
\<close>
126
+ − 179
565
+ − 180
ML %grayML %grayML\<open>fun pretty_cterm ctxt ctrm =
+ − 181
pretty_term ctxt (Thm.term_of ctrm)\<close>
126
+ − 182
565
+ − 183
text \<open>
440
+ − 184
Here the function @{ML_ind term_of in Thm} extracts the @{ML_type
+ − 185
term} from a @{ML_type cterm}. More than one @{ML_type cterm}s can be
446
+ − 186
printed again with @{ML commas in Pretty}.
565
+ − 187
\<close>
126
+ − 188
565
+ − 189
ML %grayML\<open>fun pretty_cterms ctxt ctrms =
+ − 190
Pretty.block (Pretty.commas (map (pretty_cterm ctxt) ctrms))\<close>
126
+ − 191
565
+ − 192
text \<open>
126
+ − 193
The easiest way to get the string of a theorem is to transform it
369
+ − 194
into a @{ML_type term} using the function @{ML_ind prop_of in Thm}.
565
+ − 195
\<close>
190
+ − 196
565
+ − 197
ML %grayML\<open>fun pretty_thm ctxt thm =
+ − 198
pretty_term ctxt (Thm.prop_of thm)\<close>
190
+ − 199
565
+ − 200
text \<open>
+ − 201
Theorems include schematic variables, such as \<open>?P\<close>,
+ − 202
\<open>?Q\<close> and so on. They are instantiated by Isabelle when a theorem is applied.
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 203
For example, the theorem
314
+ − 204
@{thm [source] conjI} shown below can be used for any (typable)
565
+ − 205
instantiation of \<open>?P\<close> and \<open>?Q\<close>.
190
+ − 206
572
+ − 207
@{ML_response [display, gray]
569
+ − 208
\<open>pwriteln (pretty_thm @{context} @{thm conjI})\<close>
+ − 209
\<open>\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q\<close>}
190
+ − 210
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 211
However, to improve the readability when printing theorems, we
466
+ − 212
can switch off the question marks as follows:
565
+ − 213
\<close>
126
+ − 214
565
+ − 215
ML %grayML\<open>fun pretty_thm_no_vars ctxt thm =
466
+ − 216
let
+ − 217
val ctxt' = Config.put show_question_marks false ctxt
158
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 218
in
562
+ − 219
pretty_term ctxt' (Thm.prop_of thm)
565
+ − 220
end\<close>
126
+ − 221
565
+ − 222
text \<open>
374
+ − 223
With this function, theorem @{thm [source] conjI} is now printed as follows:
190
+ − 224
572
+ − 225
@{ML_response [display, gray]
569
+ − 226
\<open>pwriteln (pretty_thm_no_vars @{context} @{thm conjI})\<close>
+ − 227
\<open>\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> P \<and> Q\<close>}
190
+ − 228
467
+ − 229
Again the functions @{ML commas} and @{ML block in Pretty} help
+ − 230
with printing more than one theorem.
565
+ − 231
\<close>
126
+ − 232
565
+ − 233
ML %grayML\<open>fun pretty_thms ctxt thms =
446
+ − 234
Pretty.block (Pretty.commas (map (pretty_thm ctxt) thms))
190
+ − 235
440
+ − 236
fun pretty_thms_no_vars ctxt thms =
565
+ − 237
Pretty.block (Pretty.commas (map (pretty_thm_no_vars ctxt) thms))\<close>
126
+ − 238
565
+ − 239
text \<open>
476
+ − 240
Printing functions for @{ML_type typ} are
565
+ − 241
\<close>
414
+ − 242
565
+ − 243
ML %grayML\<open>fun pretty_typ ctxt ty = Syntax.pretty_typ ctxt ty
446
+ − 244
fun pretty_typs ctxt tys =
565
+ − 245
Pretty.block (Pretty.commas (map (pretty_typ ctxt) tys))\<close>
414
+ − 246
565
+ − 247
text \<open>
476
+ − 248
respectively @{ML_type ctyp}
565
+ − 249
\<close>
414
+ − 250
565
+ − 251
ML %grayML\<open>fun pretty_ctyp ctxt cty = pretty_typ ctxt (Thm.typ_of cty)
446
+ − 252
fun pretty_ctyps ctxt ctys =
565
+ − 253
Pretty.block (Pretty.commas (map (pretty_ctyp ctxt) ctys))\<close>
414
+ − 254
565
+ − 255
text \<open>
370
+ − 256
\begin{readmore}
+ − 257
The simple conversion functions from Isabelle's main datatypes to
+ − 258
@{ML_type string}s are implemented in @{ML_file "Pure/Syntax/syntax.ML"}.
467
+ − 259
The configuration values that change the printing information are declared in
370
+ − 260
@{ML_file "Pure/Syntax/printer.ML"}.
+ − 261
\end{readmore}
+ − 262
414
+ − 263
Note that for printing out several ``parcels'' of information that belong
+ − 264
together, like a warning message consisting of a term and its type, you
+ − 265
should try to print these parcels together in a single string. Therefore do
+ − 266
\emph{not} print out information as
306
+ − 267
567
+ − 268
@{ML_matchresult_fake [display,gray]
569
+ − 269
\<open>pwriteln (Pretty.str "First half,");
+ − 270
pwriteln (Pretty.str "and second half.")\<close>
572
+ − 271
\<open>"First half,
+ − 272
and second half."\<close>}
306
+ − 273
+ − 274
but as a single string with appropriate formatting. For example
+ − 275
572
+ − 276
@{ML_response [display,gray]
569
+ − 277
\<open>pwriteln (Pretty.str ("First half," ^ "\\n" ^ "and second half."))\<close>
+ − 278
\<open>First half,
+ − 279
and second half.\<close>}
306
+ − 280
+ − 281
To ease this kind of string manipulations, there are a number
374
+ − 282
of library functions in Isabelle. For example, the function
+ − 283
@{ML_ind cat_lines in Library} concatenates a list of strings
+ − 284
and inserts newlines in between each element.
305
2ac9dc1a95b4
added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 285
572
+ − 286
@{ML_response [display, gray]
569
+ − 287
\<open>pwriteln (Pretty.str (cat_lines ["foo", "bar"]))\<close>
+ − 288
\<open>foo
+ − 289
bar\<close>}
306
+ − 290
414
+ − 291
Section \ref{sec:pretty} will explain the infrastructure that Isabelle
374
+ − 292
provides for more elaborate pretty printing.
350
+ − 293
+ − 294
\begin{readmore}
+ − 295
Most of the basic string functions of Isabelle are defined in
+ − 296
@{ML_file "Pure/library.ML"}.
+ − 297
\end{readmore}
565
+ − 298
\<close>
305
2ac9dc1a95b4
added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 299
2ac9dc1a95b4
added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 300
565
+ − 301
section \<open>Combinators\label{sec:combinators}\<close>
126
+ − 302
565
+ − 303
text \<open>
413
+ − 304
For beginners perhaps the most puzzling parts in the existing code of
+ − 305
Isabelle are the combinators. At first they seem to greatly obstruct the
+ − 306
comprehension of code, but after getting familiar with them and handled with
+ − 307
care, they actually ease the understanding and also the programming.
126
+ − 308
373
+ − 309
The simplest combinator is @{ML_ind I in Library}, which is just the
344
+ − 310
identity function defined as
565
+ − 311
\<close>
126
+ − 312
565
+ − 313
ML %grayML\<open>fun I x = x\<close>
126
+ − 314
565
+ − 315
text \<open>
344
+ − 316
Another simple combinator is @{ML_ind K in Library}, defined as
565
+ − 317
\<close>
126
+ − 318
565
+ − 319
ML %grayML\<open>fun K x = fn _ => x\<close>
126
+ − 320
565
+ − 321
text \<open>
+ − 322
@{ML K} ``wraps'' a function around \<open>x\<close> that ignores its argument. As a
+ − 323
result, @{ML K} defines a constant function always returning \<open>x\<close>.
126
+ − 324
344
+ − 325
The next combinator is reverse application, @{ML_ind "|>" in Basics}, defined as:
565
+ − 326
\<close>
126
+ − 327
565
+ − 328
ML %grayML\<open>fun x |> f = f x\<close>
126
+ − 329
565
+ − 330
text \<open>While just syntactic sugar for the usual function application,
126
+ − 331
the purpose of this combinator is to implement functions in a
565
+ − 332
``waterfall fashion''. Consider for example the function\<close>
126
+ − 333
565
+ − 334
ML %linenosgray\<open>fun inc_by_five x =
126
+ − 335
x |> (fn x => x + 1)
+ − 336
|> (fn x => (x, x))
+ − 337
|> fst
565
+ − 338
|> (fn x => x + 4)\<close>
126
+ − 339
565
+ − 340
text \<open>
+ − 341
which increments its argument \<open>x\<close> by 5. It does this by first
414
+ − 342
incrementing the argument by 1 (Line 2); then storing the result in a pair
+ − 343
(Line 3); taking the first component of the pair (Line 4) and finally
+ − 344
incrementing the first component by 4 (Line 5). This kind of cascading
+ − 345
manipulations of values is quite common when dealing with theories. The
+ − 346
reverse application allows you to read what happens in a top-down
+ − 347
manner. This kind of coding should be familiar, if you have been exposed to
+ − 348
Haskell's {\it do}-notation. Writing the function @{ML inc_by_five} using
+ − 349
the reverse application is much clearer than writing
565
+ − 350
\<close>
126
+ − 351
565
+ − 352
ML %grayML\<open>fun inc_by_five x = fst ((fn x => (x, x)) (x + 1)) + 4\<close>
126
+ − 353
565
+ − 354
text \<open>or\<close>
126
+ − 355
565
+ − 356
ML %grayML\<open>fun inc_by_five x =
+ − 357
((fn x => x + 4) o fst o (fn x => (x, x)) o (fn x => x + 1)) x\<close>
126
+ − 358
565
+ − 359
text \<open>and typographically more economical than\<close>
126
+ − 360
565
+ − 361
ML %grayML\<open>fun inc_by_five x =
257
+ − 362
let val y1 = x + 1
+ − 363
val y2 = (y1, y1)
+ − 364
val y3 = fst y2
+ − 365
val y4 = y3 + 4
565
+ − 366
in y4 end\<close>
126
+ − 367
565
+ − 368
text \<open>
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 369
Another reasons to avoid the let-bindings in the code above:
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 370
it is easy to get the intermediate values wrong and
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 371
the resulting code is difficult to maintain.
126
+ − 372
350
+ − 373
In Isabelle a ``real world'' example for a function written in
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 374
the waterfall fashion might be the following code:
565
+ − 375
\<close>
126
+ − 376
565
+ − 377
ML %linenosgray\<open>fun apply_fresh_args f ctxt =
193
+ − 378
f |> fastype_of
+ − 379
|> binder_types
+ − 380
|> map (pair "z")
+ − 381
|> Variable.variant_frees ctxt [f]
+ − 382
|> map Free
565
+ − 383
|> curry list_comb f\<close>
126
+ − 384
565
+ − 385
text \<open>
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 386
This function takes a term and a context as arguments. If the term is of function
569
+ − 387
type, then @{ML \<open>apply_fresh_args\<close>} returns the term with distinct variables
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 388
applied to it. For example, below three variables are applied to the term
298
+ − 389
@{term [show_types] "P::nat \<Rightarrow> int \<Rightarrow> unit \<Rightarrow> bool"}:
183
+ − 390
572
+ − 391
@{ML_response [display,gray]
569
+ − 392
\<open>let
+ − 393
val trm = @{term "P::nat \<Rightarrow> int \<Rightarrow> unit \<Rightarrow> bool"}
266
+ − 394
val ctxt = @{context}
+ − 395
in
414
+ − 396
apply_fresh_args trm ctxt
441
+ − 397
|> pretty_term ctxt
+ − 398
|> pwriteln
569
+ − 399
end\<close>
+ − 400
\<open>P z za zb\<close>}
177
+ − 401
344
+ − 402
You can read off this behaviour from how @{ML apply_fresh_args} is coded: in
+ − 403
Line 2, the function @{ML_ind fastype_of in Term} calculates the type of the
+ − 404
term; @{ML_ind binder_types in Term} in the next line produces the list of
565
+ − 405
argument types (in the case above the list \<open>[nat, int, unit]\<close>); Line
+ − 406
4 pairs up each type with the string \<open>z\<close>; the function @{ML_ind
+ − 407
variant_frees in Variable} generates for each \<open>z\<close> a unique name
+ − 408
avoiding the given \<open>f\<close>; the list of name-type pairs is turned into a
344
+ − 409
list of variable terms in Line 6, which in the last line is applied by the
414
+ − 410
function @{ML_ind list_comb in Term} to the original term. In this last step we have
344
+ − 411
to use the function @{ML_ind curry in Library}, because @{ML list_comb}
372
+ − 412
expects the function and the variables list as a pair.
374
+ − 413
414
+ − 414
Functions like @{ML apply_fresh_args} are often needed when constructing
+ − 415
terms involving fresh variables. For this the infrastructure helps
+ − 416
tremendously to avoid any name clashes. Consider for example:
252
+ − 417
572
+ − 418
@{ML_response [display,gray]
569
+ − 419
\<open>let
+ − 420
val trm = @{term "za::'a \<Rightarrow> 'b \<Rightarrow> 'c"}
266
+ − 421
val ctxt = @{context}
+ − 422
in
414
+ − 423
apply_fresh_args trm ctxt
441
+ − 424
|> pretty_term ctxt
+ − 425
|> pwriteln
569
+ − 426
end\<close>
+ − 427
\<open>za z zb\<close>}
177
+ − 428
565
+ − 429
where the \<open>za\<close> is correctly avoided.
266
+ − 430
344
+ − 431
The combinator @{ML_ind "#>" in Basics} is the reverse function composition.
+ − 432
It can be used to define the following function
565
+ − 433
\<close>
126
+ − 434
565
+ − 435
ML %grayML\<open>val inc_by_six =
374
+ − 436
(fn x => x + 1) #>
+ − 437
(fn x => x + 2) #>
565
+ − 438
(fn x => x + 3)\<close>
126
+ − 439
565
+ − 440
text \<open>
414
+ − 441
which is the function composed of first the increment-by-one function and
+ − 442
then increment-by-two, followed by increment-by-three. Again, the reverse
+ − 443
function composition allows you to read the code top-down. This combinator
+ − 444
is often used for setup functions inside the
478
+ − 445
\isacommand{setup}- or \isacommand{local\_setup}-command. These functions
+ − 446
have to be of type @{ML_type "theory -> theory"}, respectively
+ − 447
@{ML_type "local_theory -> local_theory"}. More than one such setup function
569
+ − 448
can be composed with @{ML \<open>#>\<close>}. Consider for example the following code,
478
+ − 449
where we store the theorems @{thm [source] conjI}, @{thm [source] conjunct1}
+ − 450
and @{thm [source] conjunct2} under alternative names.
565
+ − 451
\<close>
478
+ − 452
565
+ − 453
local_setup %graylinenos \<open>
478
+ − 454
let
+ − 455
fun my_note name thm = Local_Theory.note ((name, []), [thm]) #> snd
+ − 456
in
+ − 457
my_note @{binding "foo_conjI"} @{thm conjI} #>
+ − 458
my_note @{binding "bar_conjunct1"} @{thm conjunct1} #>
+ − 459
my_note @{binding "bar_conjunct2"} @{thm conjunct2}
565
+ − 460
end\<close>
478
+ − 461
565
+ − 462
text \<open>
478
+ − 463
The function @{ML_text "my_note"} in line 3 is just a wrapper for the function
567
+ − 464
@{ML_ind note in Local_Theory} in the structure @{ML_structure Local_Theory};
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 465
it stores a theorem under a name.
482
+ − 466
In lines 5 to 6 we call this function to give alternative names for the three
569
+ − 467
theorems. The point of @{ML \<open>#>\<close>} is that you can sequence such function calls.
478
+ − 468
+ − 469
The remaining combinators we describe in this section add convenience for
+ − 470
the ``waterfall method'' of writing functions. The combinator @{ML_ind tap
+ − 471
in Basics} allows you to get hold of an intermediate result (to do some
+ − 472
side-calculations or print out an intermediate result, for instance). The
+ − 473
function
565
+ − 474
\<close>
126
+ − 475
565
+ − 476
ML %linenosgray\<open>fun inc_by_three x =
126
+ − 477
x |> (fn x => x + 1)
474
+ − 478
|> tap (fn x => pwriteln (Pretty.str (@{make_string} x)))
565
+ − 479
|> (fn x => x + 2)\<close>
126
+ − 480
565
+ − 481
text \<open>
+ − 482
increments the argument first by \<open>1\<close> and then by \<open>2\<close>. In the
344
+ − 483
middle (Line 3), however, it uses @{ML tap} for printing the ``plus-one''
350
+ − 484
intermediate result. The function @{ML tap} can only be used for
+ − 485
side-calculations, because any value that is computed cannot be merged back
+ − 486
into the ``main waterfall''. To do this, you can use the next combinator.
126
+ − 487
344
+ − 488
The combinator @{ML_ind "`" in Basics} (a backtick) is similar to @{ML tap},
+ − 489
but applies a function to the value and returns the result together with the
350
+ − 490
value (as a pair). It is defined as
565
+ − 491
\<close>
350
+ − 492
565
+ − 493
ML %grayML\<open>fun `f = fn x => (f x, x)\<close>
350
+ − 494
565
+ − 495
text \<open>
350
+ − 496
An example for this combinator is the function
565
+ − 497
\<close>
126
+ − 498
565
+ − 499
ML %grayML\<open>fun inc_as_pair x =
126
+ − 500
x |> `(fn x => x + 1)
565
+ − 501
|> (fn (x, y) => (x, y + 1))\<close>
126
+ − 502
565
+ − 503
text \<open>
+ − 504
which takes \<open>x\<close> as argument, and then increments \<open>x\<close>, but also keeps
569
+ − 505
\<open>x\<close>. The intermediate result is therefore the pair @{ML \<open>(x + 1, x)\<close>
126
+ − 506
for x}. After that, the function increments the right-hand component of the
569
+ − 507
pair. So finally the result will be @{ML \<open>(x + 1, x + 1)\<close> for x}.
126
+ − 508
344
+ − 509
The combinators @{ML_ind "|>>" in Basics} and @{ML_ind "||>" in Basics} are
+ − 510
defined for functions manipulating pairs. The first applies the function to
126
+ − 511
the first component of the pair, defined as
565
+ − 512
\<close>
126
+ − 513
565
+ − 514
ML %grayML\<open>fun (x, y) |>> f = (f x, y)\<close>
126
+ − 515
565
+ − 516
text \<open>
126
+ − 517
and the second combinator to the second component, defined as
565
+ − 518
\<close>
126
+ − 519
565
+ − 520
ML %grayML\<open>fun (x, y) ||> f = (x, f y)\<close>
126
+ − 521
565
+ − 522
text \<open>
+ − 523
These two functions can, for example, be used to avoid explicit \<open>lets\<close> for
314
+ − 524
intermediate values in functions that return pairs. As an example, suppose you
308
+ − 525
want to separate a list of integers into two lists according to a
569
+ − 526
threshold. If the threshold is @{ML \<open>5\<close>}, the list @{ML \<open>[1,6,2,5,3,4]\<close>}
+ − 527
should be separated as @{ML \<open>([1,2,3,4], [6,5])\<close>}. Such a function can be
311
+ − 528
implemented as
565
+ − 529
\<close>
308
+ − 530
565
+ − 531
ML %grayML\<open>fun separate i [] = ([], [])
308
+ − 532
| separate i (x::xs) =
+ − 533
let
+ − 534
val (los, grs) = separate i xs
+ − 535
in
+ − 536
if i <= x then (los, x::grs) else (x::los, grs)
565
+ − 537
end\<close>
308
+ − 538
565
+ − 539
text \<open>
350
+ − 540
where the return value of the recursive call is bound explicitly to
569
+ − 541
the pair @{ML \<open>(los, grs)\<close> for los grs}. However, this function
414
+ − 542
can be implemented more concisely as
565
+ − 543
\<close>
308
+ − 544
565
+ − 545
ML %grayML\<open>fun separate _ [] = ([], [])
308
+ − 546
| separate i (x::xs) =
+ − 547
if i <= x
+ − 548
then separate i xs ||> cons x
565
+ − 549
else separate i xs |>> cons x\<close>
308
+ − 550
565
+ − 551
text \<open>
+ − 552
avoiding the explicit \<open>let\<close>. While in this example the gain in
314
+ − 553
conciseness is only small, in more complicated situations the benefit of
565
+ − 554
avoiding \<open>lets\<close> can be substantial.
308
+ − 555
344
+ − 556
With the combinator @{ML_ind "|->" in Basics} you can re-combine the
+ − 557
elements from a pair. This combinator is defined as
565
+ − 558
\<close>
126
+ − 559
565
+ − 560
ML %grayML\<open>fun (x, y) |-> f = f x y\<close>
126
+ − 561
565
+ − 562
text \<open>
215
+ − 563
and can be used to write the following roundabout version
565
+ − 564
of the \<open>double\<close> function:
+ − 565
\<close>
126
+ − 566
565
+ − 567
ML %grayML\<open>fun double x =
126
+ − 568
x |> (fn x => (x, x))
565
+ − 569
|-> (fn x => fn y => x + y)\<close>
126
+ − 570
565
+ − 571
text \<open>
564
+ − 572
The combinator @{ML_ind ||>> in Basics} plays a central role whenever your
344
+ − 573
task is to update a theory and the update also produces a side-result (for
+ − 574
example a theorem). Functions for such tasks return a pair whose second
+ − 575
component is the theory and the fist component is the side-result. Using
+ − 576
@{ML ||>>}, you can do conveniently the update and also
+ − 577
accumulate the side-results. Consider the following simple function.
565
+ − 578
\<close>
215
+ − 579
565
+ − 580
ML %linenosgray\<open>fun acc_incs x =
215
+ − 581
x |> (fn x => ("", x))
+ − 582
||>> (fn x => (x, x + 1))
+ − 583
||>> (fn x => (x, x + 1))
565
+ − 584
||>> (fn x => (x, x + 1))\<close>
215
+ − 585
565
+ − 586
text \<open>
215
+ − 587
The purpose of Line 2 is to just pair up the argument with a dummy value (since
344
+ − 588
@{ML ||>>} operates on pairs). Each of the next three lines just increment
280
+ − 589
the value by one, but also nest the intermediate results to the left. For example
215
+ − 590
567
+ − 591
@{ML_matchresult [display,gray]
569
+ − 592
\<open>acc_incs 1\<close>
+ − 593
\<open>(((("", 1), 2), 3), 4)\<close>}
215
+ − 594
+ − 595
You can continue this chain with:
+ − 596
567
+ − 597
@{ML_matchresult [display,gray]
569
+ − 598
\<open>acc_incs 1 ||>> (fn x => (x, x + 2))\<close>
+ − 599
\<open>((((("", 1), 2), 3), 4), 6)\<close>}
215
+ − 600
483
+ − 601
An example where this combinator is useful is as follows
+ − 602
572
+ − 603
@{ML_matchresult [display, gray]
569
+ − 604
\<open>let
483
+ − 605
val ((names1, names2), _) =
+ − 606
@{context}
569
+ − 607
|> Variable.variant_fixes (replicate 4 "x")
+ − 608
||>> Variable.variant_fixes (replicate 5 "x")
483
+ − 609
in
+ − 610
(names1, names2)
569
+ − 611
end\<close>
+ − 612
\<open>(["x", "xa", "xb", "xc"], ["xd", "xe", "xf", "xg", "xh"])\<close>}
483
+ − 613
569
+ − 614
Its purpose is to create nine variants of the string @{ML \<open>"x"\<close>} so
483
+ − 615
that no variant will clash with another. Suppose for some reason we want
+ − 616
to bind four variants to the lists @{ML_text "name1"} and the
+ − 617
rest to @{ML_text "name2"}. In order to obtain non-clashing
+ − 618
variants we have to thread the context through the function calls
+ − 619
(the context records which variants have been previously created).
569
+ − 620
For the first call we can use @{ML \<open>|>\<close>}, but in the
483
+ − 621
second and any further call to @{ML_ind variant_fixes in Variable} we
569
+ − 622
have to use @{ML \<open>||>>\<close>} in order to account for the result(s)
483
+ − 623
obtained by previous calls.
+ − 624
479
+ − 625
A more realistic example for this combinator is the following code
565
+ − 626
\<close>
478
+ − 627
565
+ − 628
ML %grayML\<open>val ((([one_def], [two_def]), [three_def]), ctxt') =
478
+ − 629
@{context}
562
+ − 630
|> Local_Defs.define [((@{binding "One"}, NoSyn), (Binding.empty_atts, @{term "1::nat"}))]
+ − 631
||>> Local_Defs.define [((@{binding "Two"}, NoSyn), (Binding.empty_atts,@{term "2::nat"}))]
565
+ − 632
||>> Local_Defs.define [((@{binding "Three"}, NoSyn), (Binding.empty_atts,@{term "3::nat"}))]\<close>
478
+ − 633
565
+ − 634
text \<open>
496
+ − 635
where we make three definitions, namely @{term "One \<equiv> 1::nat"}, @{term "Two \<equiv> 2::nat"}
+ − 636
and @{term "Three \<equiv> 3::nat"}. The point of this code is that we augment the initial
479
+ − 637
context with the definitions. The result we are interested in is the
+ − 638
augmented context, that is @{ML_text "ctxt'"}, but also the side-results containing
562
+ − 639
information about the definitions---the function @{ML_ind define in Local_Defs} returns
479
+ − 640
both as pairs. We can use this information for example to print out the definiens and
+ − 641
the theorem corresponding to the definitions. For example for the first definition:
+ − 642
572
+ − 643
@{ML_response [display, gray]
569
+ − 644
\<open>let
562
+ − 645
val (one_trm, (_, one_thm)) = one_def
479
+ − 646
in
572
+ − 647
(pwriteln (pretty_term ctxt' one_trm),
+ − 648
pwriteln (pretty_thm ctxt' one_thm))
569
+ − 649
end\<close>
+ − 650
\<open>One
+ − 651
One \<equiv> 1\<close>}
+ − 652
Recall that @{ML \<open>|>\<close>} is the reverse function application. Recall also that
+ − 653
the related reverse function composition is @{ML \<open>#>\<close>}. In fact all the
+ − 654
combinators @{ML \<open>|->\<close>}, @{ML \<open>|>>\<close>} , @{ML \<open>||>\<close>} and @{ML \<open>||>>\<close>}
344
+ − 655
described above have related combinators for function composition, namely
+ − 656
@{ML_ind "#->" in Basics}, @{ML_ind "#>>" in Basics}, @{ML_ind "##>" in
569
+ − 657
Basics} and @{ML_ind "##>>" in Basics}. Using @{ML \<open>#->\<close>}, for example, the
565
+ − 658
function \<open>double\<close> can also be written as:
+ − 659
\<close>
126
+ − 660
565
+ − 661
ML %grayML\<open>val double =
502
+ − 662
(fn x => (x, x)) #->
565
+ − 663
(fn x => fn y => x + y)\<close>
126
+ − 664
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 665
565
+ − 666
text \<open>
314
+ − 667
When using combinators for writing functions in waterfall fashion, it is
311
+ − 668
sometimes necessary to do some ``plumbing'' in order to fit functions
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 669
together. We have already seen such plumbing in the function @{ML
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 670
apply_fresh_args}, where @{ML curry} is needed for making the function @{ML
569
+ − 671
list_comb}, which works over pairs, to fit with the combinator @{ML \<open>|>\<close>}. Such
414
+ − 672
plumbing is also needed in situations where a function operates over lists,
325
+ − 673
but one calculates only with a single element. An example is the function
350
+ − 674
@{ML_ind check_terms in Syntax}, whose purpose is to simultaneously type-check
+ − 675
a list of terms. Consider the code:
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 676
572
+ − 677
@{ML_response [display, gray]
569
+ − 678
\<open>let
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 679
val ctxt = @{context}
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 680
in
569
+ − 681
map (Syntax.parse_term ctxt) ["m + n", "m * n", "m - (n::nat)"]
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 682
|> Syntax.check_terms ctxt
441
+ − 683
|> pretty_terms ctxt
+ − 684
|> pwriteln
569
+ − 685
end\<close>
+ − 686
\<open>m + n, m * n, m - n\<close>}
565
+ − 687
\<close>
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 688
565
+ − 689
text \<open>
372
+ − 690
In this example we obtain three terms (using the function
565
+ − 691
@{ML_ind parse_term in Syntax}) whose variables \<open>m\<close> and \<open>n\<close>
372
+ − 692
are of type @{typ "nat"}. If you have only a single term, then @{ML
324
+ − 693
check_terms in Syntax} needs plumbing. This can be done with the function
372
+ − 694
@{ML_ind singleton in Library}.\footnote{There is already a function @{ML check_term in
+ − 695
Syntax} in the file @{ML_file "Pure/Syntax/syntax.ML"} that is implemented
+ − 696
in terms of @{ML singleton} and @{ML check_terms in Syntax}.} For example
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 697
572
+ − 698
@{ML_response [display, gray, linenos]
569
+ − 699
\<open>let
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 700
val ctxt = @{context}
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 701
in
569
+ − 702
Syntax.parse_term ctxt "m - (n::nat)"
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 703
|> singleton (Syntax.check_terms ctxt)
441
+ − 704
|> pretty_term ctxt
+ − 705
|> pwriteln
569
+ − 706
end\<close>
+ − 707
\<open>m - n\<close>}
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 708
372
+ − 709
where in Line 5, the function operating over lists fits with the
+ − 710
single term generated in Line 4.
+ − 711
127
+ − 712
\begin{readmore}
196
+ − 713
The most frequently used combinators are defined in the files @{ML_file
+ − 714
"Pure/library.ML"}
127
+ − 715
and @{ML_file "Pure/General/basics.ML"}. Also \isccite{sec:ML-linear-trans}
+ − 716
contains further information about combinators.
+ − 717
\end{readmore}
310
007922777ff1
added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 718
421
+ − 719
\begin{exercise}
569
+ − 720
Find out what the combinator @{ML \<open>K I\<close>} does.
421
+ − 721
\end{exercise}
565
+ − 722
\<close>
15
+ − 723
10
+ − 724
565
+ − 725
section \<open>ML-Antiquotations\label{sec:antiquote}\<close>
2
+ − 726
565
+ − 727
text \<open>
372
+ − 728
Recall from Section \ref{sec:include} that code in Isabelle is always
+ − 729
embedded in a theory. The main advantage of this is that the code can
+ − 730
contain references to entities defined on the logical level of Isabelle. By
+ − 731
this we mean references to definitions, theorems, terms and so on. These
+ − 732
reference are realised in Isabelle with ML-antiquotations, often just called
+ − 733
antiquotations.\footnote{Note that there are two kinds of antiquotations in
+ − 734
Isabelle, which have very different purposes and infrastructures. The first
+ − 735
kind, described in this section, are \emph{\index*{ML-antiquotation}}. They
+ − 736
are used to refer to entities (like terms, types etc) from Isabelle's logic
+ − 737
layer inside ML-code. The other kind of antiquotations are
+ − 738
\emph{document}\index{document antiquotation} antiquotations. They are used
+ − 739
only in the text parts of Isabelle and their purpose is to print logical
+ − 740
entities inside \LaTeX-documents. Document antiquotations are part of the
+ − 741
user level and therefore we are not interested in them in this Tutorial,
+ − 742
except in Appendix \ref{rec:docantiquotations} where we show how to
+ − 743
implement your own document antiquotations.} Syntactically antiquotations
565
+ − 744
are indicated by the @{ML_text @}-sign followed by text wrapped in \<open>{\<dots>}\<close>. For example, one can print out the name of the current theory with
372
+ − 745
the code
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 746
569
+ − 747
@{ML_matchresult [display,gray] \<open>Context.theory_name @{theory}\<close> \<open>"First_Steps"\<close>}
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 748
565
+ − 749
where \<open>@{theory}\<close> is an antiquotation that is substituted with the
49
+ − 750
current theory (remember that we assumed we are inside the theory
565
+ − 751
\<open>First_Steps\<close>). The name of this theory can be extracted using
344
+ − 752
the function @{ML_ind theory_name in Context}.
5
+ − 753
89
+ − 754
Note, however, that antiquotations are statically linked, that is their value is
329
+ − 755
determined at ``compile-time'', not at ``run-time''. For example the function
565
+ − 756
\<close>
5
+ − 757
565
+ − 758
ML %grayML\<open>fun not_current_thyname () = Context.theory_name @{theory}\<close>
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 759
565
+ − 760
text \<open>
89
+ − 761
does \emph{not} return the name of the current theory, if it is run in a
5
+ − 762
different theory. Instead, the code above defines the constant function
441
+ − 763
that always returns the string @{text [quotes] "First_Steps"}, no matter where the
565
+ − 764
function is called. Operationally speaking, the antiquotation \<open>@{theory}\<close> is
5
+ − 765
\emph{not} replaced with code that will look up the current theory in
+ − 766
some data structure and return it. Instead, it is literally
414
+ − 767
replaced with the value representing the theory.
+ − 768
565
+ − 769
Another important antiquotation is \<open>@{context}\<close>. (What the
414
+ − 770
difference between a theory and a context is will be described in Chapter
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 771
\ref{chp:advanced}.) A context is for example needed to use the
475
+ − 772
function @{ML print_abbrevs in Proof_Context} that list of all currently
557
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 773
defined abbreviations. For example
2
+ − 774
567
+ − 775
@{ML_matchresult_fake [display, gray]
569
+ − 776
\<open>Proof_Context.print_abbrevs false @{context}\<close>
+ − 777
\<open>\<dots>
572
+ − 778
INTER \<equiv> INFIMUM
414
+ − 779
Inter \<equiv> Inf
569
+ − 780
\<dots>\<close>}
414
+ − 781
557
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 782
The precise output of course depends on the abbreviations that are
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 783
currently defined; this can change over time.
414
+ − 784
You can also use antiquotations to refer to proved theorems:
565
+ − 785
\<open>@{thm \<dots>}\<close> for a single theorem
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 786
572
+ − 787
@{ML_response [display,gray] \<open>@{thm allI}\<close> \<open>(\<And>x. ?P x) \<Longrightarrow> \<forall>x. ?P x\<close>}
75
+ − 788
565
+ − 789
and \<open>@{thms \<dots>}\<close> for more than one
132
+ − 790
572
+ − 791
+ − 792
@{ML_response [display,gray]
569
+ − 793
\<open>@{thms conj_ac}\<close>
572
+ − 794
\<open>["(?P \<and> ?Q) = (?Q \<and> ?P)",
+ − 795
"(?P \<and> ?Q \<and> ?R) = (?Q \<and> ?P \<and> ?R)",
+ − 796
"((?P \<and> ?Q) \<and> ?R) = (?P \<and> ?Q \<and> ?R)"]\<close>}
132
+ − 797
414
+ − 798
The thm-antiquotations can also be used for manipulating theorems. For
474
+ − 799
example, if you need the version of the theorem @{thm [source] refl} that
414
+ − 800
has a meta-equality instead of an equality, you can write
+ − 801
572
+ − 802
@{ML_response [display,gray]
569
+ − 803
\<open>@{thm refl[THEN eq_reflection]}\<close>
+ − 804
\<open>?x \<equiv> ?x\<close>}
414
+ − 805
292
+ − 806
The point of these antiquotations is that referring to theorems in this way
+ − 807
makes your code independent from what theorems the user might have stored
+ − 808
under this name (this becomes especially important when you deal with
329
+ − 809
theorem lists; see Section \ref{sec:storing}).
292
+ − 810
565
+ − 811
It is also possible to prove lemmas with the antiquotation \<open>@{lemma \<dots> by \<dots>}\<close>
+ − 812
whose first argument is a statement (possibly many of them separated by \<open>and\<close>)
375
+ − 813
and the second is a proof. For example
565
+ − 814
\<close>
375
+ − 815
565
+ − 816
ML %grayML\<open>val foo_thms = @{lemma "True" and "False \<Longrightarrow> P" by simp_all}\<close>
375
+ − 817
565
+ − 818
text \<open>
377
+ − 819
The result can be printed out as follows.
375
+ − 820
572
+ − 821
@{ML_response [gray,display]
569
+ − 822
\<open>foo_thms |> pretty_thms_no_vars @{context}
+ − 823
|> pwriteln\<close>
+ − 824
\<open>True, False \<Longrightarrow> P\<close>}
375
+ − 825
292
+ − 826
You can also refer to the current simpset via an antiquotation. To illustrate
+ − 827
this we implement the function that extracts the theorem names stored in a
+ − 828
simpset.
565
+ − 829
\<close>
75
+ − 830
565
+ − 831
ML %grayML\<open>fun get_thm_names_from_ss ctxt =
131
+ − 832
let
544
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 833
val simpset = Raw_Simplifier.simpset_of ctxt
458
+ − 834
val {simps,...} = Raw_Simplifier.dest_ss simpset
70
+ − 835
in
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
+ − 836
map #1 simps
565
+ − 837
end\<close>
54
+ − 838
565
+ − 839
text \<open>
458
+ − 840
The function @{ML_ind dest_ss in Raw_Simplifier} returns a record containing all
414
+ − 841
information stored in the simpset, but here we are only interested in the names of the
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 842
simp-rules. Now you can feed in the current simpset into this function.
544
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 843
The current simpset can be referred to using @{ML_ind simpset_of in Raw_Simplifier}.
81
+ − 844
572
+ − 845
@{ML_response [display,gray]
569
+ − 846
\<open>get_thm_names_from_ss @{context}\<close>
572
+ − 847
\<open>["Euclidean_Division.euclidean_size_int_def",\<dots>\<close>}
10
+ − 848
196
+ − 849
Again, this way of referencing simpsets makes you independent from additions
350
+ − 850
of lemmas to the simpset by the user, which can potentially cause loops in your
292
+ − 851
code.
156
+ − 852
292
+ − 853
It is also possible to define your own antiquotations. But you should
315
+ − 854
exercise care when introducing new ones, as they can also make your code
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 855
difficult to read. In the next chapter we describe how to construct
565
+ − 856
terms with the (built-in) antiquotation \<open>@{term \<dots>}\<close>. A restriction
372
+ − 857
of this antiquotation is that it does not allow you to use schematic
+ − 858
variables in terms. If you want to have an antiquotation that does not have
323
+ − 859
this restriction, you can implement your own using the function @{ML_ind
567
+ − 860
inline in ML_Antiquotation} from the structure @{ML_structure ML_Antiquotation}. The code
565
+ − 861
for the antiquotation \<open>term_pat\<close> is as follows.
+ − 862
\<close>
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 863
565
+ − 864
ML %linenosgray\<open>val term_pat_setup =
471
+ − 865
let
562
+ − 866
val parser = Args.context -- Scan.lift Args.embedded_inner_syntax
325
+ − 867
+ − 868
fun term_pat (ctxt, str) =
475
+ − 869
str |> Proof_Context.read_term_pattern ctxt
264
+ − 870
|> ML_Syntax.print_term
325
+ − 871
|> ML_Syntax.atomic
+ − 872
in
554
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 873
ML_Antiquotation.inline @{binding "term_pat"} (parser >> term_pat)
565
+ − 874
end\<close>
263
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 875
565
+ − 876
text \<open>
474
+ − 877
To use it you also have to install it using \isacommand{setup} like so
565
+ − 878
\<close>
474
+ − 879
565
+ − 880
setup %gray \<open>term_pat_setup\<close>
471
+ − 881
565
+ − 882
text \<open>
575
+ − 883
The parser in Line 3 provides us with a context and a string; this string is
324
+ − 884
transformed into a term using the function @{ML_ind read_term_pattern in
575
+ − 885
Proof_Context} (Line 6); the next two lines transform the term into a string
372
+ − 886
so that the ML-system can understand it. (All these functions will be explained
+ − 887
in more detail in later sections.) An example for this antiquotation is:
292
+ − 888
572
+ − 889
@{ML_matchresult [display,gray]
569
+ − 890
\<open>@{term_pat "Suc (?x::nat)"}\<close>
572
+ − 891
\<open>Const ("Nat.Suc", _) $ Var (("x", 0), _)\<close>}
292
+ − 892
565
+ − 893
which shows the internal representation of the term \<open>Suc ?x\<close>. Similarly
474
+ − 894
we can write an antiquotation for type patterns. Its code is
565
+ − 895
\<close>
377
+ − 896
565
+ − 897
ML %grayML\<open>val type_pat_setup =
471
+ − 898
let
562
+ − 899
val parser = Args.context -- Scan.lift Args.embedded_inner_syntax
298
+ − 900
377
+ − 901
fun typ_pat (ctxt, str) =
503
+ − 902
let
+ − 903
val ctxt' = Proof_Context.set_mode Proof_Context.mode_schematic ctxt
+ − 904
in
+ − 905
str |> Syntax.read_typ ctxt'
+ − 906
|> ML_Syntax.print_typ
+ − 907
|> ML_Syntax.atomic
+ − 908
end
377
+ − 909
in
554
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 910
ML_Antiquotation.inline @{binding "typ_pat"} (parser >> typ_pat)
565
+ − 911
end\<close>
377
+ − 912
565
+ − 913
text \<open>
474
+ − 914
which can be installed with
565
+ − 915
\<close>
474
+ − 916
565
+ − 917
setup %gray \<open>type_pat_setup\<close>
471
+ − 918
565
+ − 919
text \<open>
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 920
However, a word of warning is in order: new antiquotations should be introduced only after
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 921
careful deliberations. They can potentially make your code harder, rather than easier, to read.
474
+ − 922
562
+ − 923
\begin{readmore} @{url "Norbert/ML/ml_antiquotation.ML"}
554
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 924
The files @{ML_file "Pure/ML/ml_antiquotation.ML"} and @{ML_file "Pure/ML/ml_antiquotations.ML"}
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 925
contain the infrastructure and definitions for most antiquotations. Most of the basic operations
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 926
on ML-syntax are implemented in @{ML_file "Pure/ML/ml_syntax.ML"}.
263
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 927
\end{readmore}
565
+ − 928
\<close>
323
+ − 929
565
+ − 930
section \<open>Storing Data in Isabelle\label{sec:storing}\<close>
292
+ − 931
565
+ − 932
text \<open>
327
+ − 933
Isabelle provides mechanisms for storing (and retrieving) arbitrary
+ − 934
data. Before we delve into the details, let us digress a bit. Conventional
350
+ − 935
wisdom has it that the type-system of ML ensures that an
+ − 936
@{ML_type "'a list"}, say, can only hold elements of the same type, namely
467
+ − 937
@{ML_type "'a"} (or whatever is substitued for it). Despite this common
+ − 938
wisdom, however, it is possible to implement a
327
+ − 939
universal type in ML, although by some arguably accidental features of ML.
+ − 940
This universal type can be used to store data of different type into a single list.
350
+ − 941
In fact, it allows one to inject and to project data of \emph{arbitrary} type. This is
327
+ − 942
in contrast to datatypes, which only allow injection and projection of data for
372
+ − 943
some \emph{fixed} collection of types. In light of the conventional wisdom cited
327
+ − 944
above it is important to keep in mind that the universal type does not
+ − 945
destroy type-safety of ML: storing and accessing the data can only be done
474
+ − 946
in a type-safe manner...though run-time checks are needed for that.
323
+ − 947
327
+ − 948
\begin{readmore}
562
+ − 949
In ML the universal type is implemented as the type @{ML_type
+ − 950
Universal.universal}.
327
+ − 951
\end{readmore}
+ − 952
+ − 953
We will show the usage of the universal type by storing an integer and
+ − 954
a boolean into a single list. Let us first define injection and projection
350
+ − 955
functions for booleans and integers into and from the type @{ML_type Universal.universal}.
565
+ − 956
\<close>
323
+ − 957
565
+ − 958
ML %grayML\<open>local
325
+ − 959
val fn_int = Universal.tag () : int Universal.tag
+ − 960
val fn_bool = Universal.tag () : bool Universal.tag
+ − 961
in
+ − 962
val inject_int = Universal.tagInject fn_int;
+ − 963
val inject_bool = Universal.tagInject fn_bool;
+ − 964
val project_int = Universal.tagProject fn_int;
+ − 965
val project_bool = Universal.tagProject fn_bool
565
+ − 966
end\<close>
298
+ − 967
565
+ − 968
text \<open>
327
+ − 969
Using the injection functions, we can inject the integer @{ML_text "13"}
330
+ − 970
and the boolean value @{ML_text "true"} into @{ML_type Universal.universal}, and
327
+ − 971
then store them in a @{ML_type "Universal.universal list"} as follows:
565
+ − 972
\<close>
323
+ − 973
565
+ − 974
ML %grayML\<open>val foo_list =
327
+ − 975
let
474
+ − 976
val thirteen = inject_int 13
327
+ − 977
val truth_val = inject_bool true
+ − 978
in
+ − 979
[thirteen, truth_val]
565
+ − 980
end\<close>
327
+ − 981
565
+ − 982
text \<open>
372
+ − 983
The data can be retrieved with the projection functions defined above.
327
+ − 984
572
+ − 985
@{ML_response [display, gray]
+ − 986
\<open>(project_int (nth foo_list 0),
+ − 987
project_bool (nth foo_list 1))\<close>
+ − 988
\<open>13,
569
+ − 989
true\<close>}
327
+ − 990
+ − 991
Notice that we access the integer as an integer and the boolean as
+ − 992
a boolean. If we attempt to access the integer as a boolean, then we get
+ − 993
a runtime error.
557
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 994
572
+ − 995
@{ML_response [display, gray]
569
+ − 996
\<open>project_bool (nth foo_list 0)\<close>
572
+ − 997
\<open>exception Match raised\<close>}
327
+ − 998
+ − 999
This runtime error is the reason why ML is still type-sound despite
+ − 1000
containing a universal type.
+ − 1001
328
+ − 1002
Now, Isabelle heavily uses this mechanism for storing all sorts of
+ − 1003
data: theorem lists, simpsets, facts etc. Roughly speaking, there are two
350
+ − 1004
places where data can be stored in Isabelle: in \emph{theories} and in \emph{proof
372
+ − 1005
contexts}. Data such as simpsets are ``global'' and therefore need to be stored
+ − 1006
in a theory (simpsets need to be maintained across proofs and even across
+ − 1007
theories). On the other hand, data such as facts change inside a proof and
328
+ − 1008
are only relevant to the proof at hand. Therefore such data needs to be
372
+ − 1009
maintained inside a proof context, which represents ``local'' data.
467
+ − 1010
You can think of a theory as the ``longterm memory'' of Isabelle (nothing will
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1011
be deleted from it), and a proof-context as a ``shortterm memory'' (it
467
+ − 1012
changes according to what is needed at the time).
327
+ − 1013
328
+ − 1014
For theories and proof contexts there are, respectively, the functors
567
+ − 1015
@{ML_functor_ind Theory_Data} and @{ML_functor_ind Proof_Data} that help
372
+ − 1016
with the data storage. Below we show how to implement a table in which you
328
+ − 1017
can store theorems and look them up according to a string key. The
350
+ − 1018
intention in this example is to be able to look up introduction rules for logical
327
+ − 1019
connectives. Such a table might be useful in an automatic proof procedure
350
+ − 1020
and therefore it makes sense to store this data inside a theory.
567
+ − 1021
Consequently we use the functor @{ML_functor Theory_Data}.
350
+ − 1022
The code for the table is:
565
+ − 1023
\<close>
323
+ − 1024
565
+ − 1025
ML %linenosgray\<open>structure Data = Theory_Data
327
+ − 1026
(type T = thm Symtab.table
+ − 1027
val empty = Symtab.empty
+ − 1028
val extend = I
565
+ − 1029
val merge = Symtab.merge (K true))\<close>
327
+ − 1030
565
+ − 1031
text \<open>
327
+ − 1032
In order to store data in a theory, we have to specify the type of the data
350
+ − 1033
(Line 2). In this case we specify the type @{ML_type "thm Symtab.table"},
+ − 1034
which stands for a table in which @{ML_type string}s can be looked up
+ − 1035
producing an associated @{ML_type thm}. We also have to specify four
+ − 1036
functions to use this functor: namely how to initialise the data storage
385
+ − 1037
(Line 3), how to extend it (Line 4) and how two
+ − 1038
tables should be merged (Line 5). These functions correspond roughly to the
350
+ − 1039
operations performed on theories and we just give some sensible
372
+ − 1040
defaults.\footnote{\bf FIXME: Say more about the
350
+ − 1041
assumptions of these operations.} The result structure @{ML_text Data}
327
+ − 1042
contains functions for accessing the table (@{ML Data.get}) and for updating
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1043
it (@{ML Data.map}). There is also the function @{ML Data.put}, but it is
385
+ − 1044
not relevant here. Below we define two
350
+ − 1045
auxiliary functions, which help us with accessing the table.
565
+ − 1046
\<close>
327
+ − 1047
565
+ − 1048
ML %grayML\<open>val lookup = Symtab.lookup o Data.get
+ − 1049
fun update k v = Data.map (Symtab.update (k, v))\<close>
327
+ − 1050
565
+ − 1051
text \<open>
327
+ − 1052
Since we want to store introduction rules associated with their
+ − 1053
logical connective, we can fill the table as follows.
565
+ − 1054
\<close>
327
+ − 1055
565
+ − 1056
setup %gray \<open>
450
+ − 1057
update "conj" @{thm conjI} #>
+ − 1058
update "imp" @{thm impI} #>
+ − 1059
update "all" @{thm allI}
565
+ − 1060
\<close>
326
+ − 1061
565
+ − 1062
text \<open>
328
+ − 1063
The use of the command \isacommand{setup} makes sure the table in the
350
+ − 1064
\emph{current} theory is updated (this is explained further in
559
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 1065
Section~\ref{sec:theories}). The lookup can now be performed as follows.
327
+ − 1066
572
+ − 1067
@{ML_response [display, gray]
569
+ − 1068
\<open>lookup @{theory} "conj"\<close>
+ − 1069
\<open>SOME "\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q"\<close>}
327
+ − 1070
+ − 1071
An important point to note is that these tables (and data in general)
+ − 1072
need to be treated in a purely functional fashion. Although
+ − 1073
we can update the table as follows
565
+ − 1074
\<close>
327
+ − 1075
565
+ − 1076
setup %gray \<open>update "conj" @{thm TrueI}\<close>
327
+ − 1077
572
+ − 1078
565
+ − 1079
text \<open>
350
+ − 1080
and accordingly, @{ML lookup} now produces the introduction rule for @{term "True"}
327
+ − 1081
572
+ − 1082
@{ML_response [display, gray]
569
+ − 1083
\<open>lookup @{theory} "conj"\<close>
+ − 1084
\<open>SOME "True"\<close>}
327
+ − 1085
+ − 1086
there are no references involved. This is one of the most fundamental
350
+ − 1087
coding conventions for programming in Isabelle. References
328
+ − 1088
interfere with the multithreaded execution model of Isabelle and also
350
+ − 1089
defeat its undo-mechanism. To see the latter, consider the
328
+ − 1090
following data container where we maintain a reference to a list of
+ − 1091
integers.
565
+ − 1092
\<close>
327
+ − 1093
565
+ − 1094
ML %grayML\<open>structure WrongRefData = Theory_Data
328
+ − 1095
(type T = (int list) Unsynchronized.ref
+ − 1096
val empty = Unsynchronized.ref []
+ − 1097
val extend = I
565
+ − 1098
val merge = fst)\<close>
327
+ − 1099
565
+ − 1100
text \<open>
328
+ − 1101
We initialise the reference with the empty list. Consequently a first
569
+ − 1102
lookup produces @{ML \<open>ref []\<close> in Unsynchronized}.
328
+ − 1103
567
+ − 1104
@{ML_matchresult_fake [display,gray]
569
+ − 1105
\<open>WrongRefData.get @{theory}\<close>
+ − 1106
\<open>ref []\<close>}
328
+ − 1107
329
+ − 1108
For updating the reference we use the following function
565
+ − 1109
\<close>
328
+ − 1110
565
+ − 1111
ML %grayML\<open>fun ref_update n = WrongRefData.map
+ − 1112
(fn r => let val _ = r := n::(!r) in r end)\<close>
328
+ − 1113
565
+ − 1114
text \<open>
329
+ − 1115
which takes an integer and adds it to the content of the reference.
350
+ − 1116
As before, we update the reference with the command
328
+ − 1117
\isacommand{setup}.
565
+ − 1118
\<close>
327
+ − 1119
565
+ − 1120
setup %gray \<open>ref_update 1\<close>
327
+ − 1121
565
+ − 1122
text \<open>
328
+ − 1123
A lookup in the current theory gives then the expected list
569
+ − 1124
@{ML \<open>ref [1]\<close> in Unsynchronized}.
328
+ − 1125
567
+ − 1126
@{ML_matchresult_fake [display,gray]
569
+ − 1127
\<open>WrongRefData.get @{theory}\<close>
+ − 1128
\<open>ref [1]\<close>}
328
+ − 1129
347
+ − 1130
So far everything is as expected. But, the trouble starts if we attempt to
350
+ − 1131
backtrack to the ``point'' before the \isacommand{setup}-command. There, we
347
+ − 1132
would expect that the list is empty again. But since it is stored in a
+ − 1133
reference, Isabelle has no control over it. So it is not empty, but still
569
+ − 1134
@{ML \<open>ref [1]\<close> in Unsynchronized}. Adding to the trouble, if we execute the
+ − 1135
\isacommand{setup}-command again, we do not obtain @{ML \<open>ref [1]\<close> in
347
+ − 1136
Unsynchronized}, but
327
+ − 1137
567
+ − 1138
@{ML_matchresult_fake [display,gray]
569
+ − 1139
\<open>WrongRefData.get @{theory}\<close>
+ − 1140
\<open>ref [1, 1]\<close>}
328
+ − 1141
474
+ − 1142
Now imagine how often you go backwards and forwards in your proof
+ − 1143
scripts.\footnote{The same problem can be triggered in the Jedit GUI by
+ − 1144
making the parser to go over and over again over the \isacommand{setup} command.}
328
+ − 1145
By using references in Isabelle code, you are bound to cause all
329
+ − 1146
hell to break loose. Therefore observe the coding convention:
328
+ − 1147
Do not use references for storing data!
327
+ − 1148
328
+ − 1149
\begin{readmore}
+ − 1150
The functors for data storage are defined in @{ML_file "Pure/context.ML"}.
+ − 1151
Isabelle contains implementations of several container data structures,
+ − 1152
including association lists in @{ML_file "Pure/General/alist.ML"},
347
+ − 1153
directed graphs in @{ML_file "Pure/General/graph.ML"}, and
328
+ − 1154
tables and symtables in @{ML_file "Pure/General/table.ML"}.
+ − 1155
\end{readmore}
327
+ − 1156
350
+ − 1157
Storing data in a proof context is done in a similar fashion. As mentioned
567
+ − 1158
before, the corresponding functor is @{ML_functor_ind Proof_Data}. With the
350
+ − 1159
following code we can store a list of terms in a proof context.
565
+ − 1160
\<close>
328
+ − 1161
565
+ − 1162
ML %grayML\<open>structure Data = Proof_Data
328
+ − 1163
(type T = term list
565
+ − 1164
fun init _ = [])\<close>
328
+ − 1165
565
+ − 1166
text \<open>
414
+ − 1167
The init-function we have to specify must produce a list for when a context
350
+ − 1168
is initialised (possibly taking the theory into account from which the
372
+ − 1169
context is derived). We choose here to just return the empty list. Next
328
+ − 1170
we define two auxiliary functions for updating the list with a given
+ − 1171
term and printing the list.
565
+ − 1172
\<close>
328
+ − 1173
565
+ − 1174
ML %grayML\<open>fun update trm = Data.map (fn trms => trm::trms)
328
+ − 1175
+ − 1176
fun print ctxt =
+ − 1177
case (Data.get ctxt) of
474
+ − 1178
[] => pwriteln (Pretty.str "Empty!")
565
+ − 1179
| trms => pwriteln (pretty_terms ctxt trms)\<close>
328
+ − 1180
565
+ − 1181
text \<open>
330
+ − 1182
Next we start with the context generated by the antiquotation
565
+ − 1183
\<open>@{context}\<close> and update it in various ways.
328
+ − 1184
567
+ − 1185
@{ML_matchresult_fake [display,gray]
569
+ − 1186
\<open>let
347
+ − 1187
val ctxt0 = @{context}
569
+ − 1188
val ctxt1 = ctxt0 |> update @{term "False"}
+ − 1189
|> update @{term "True \<and> True"}
+ − 1190
val ctxt2 = ctxt0 |> update @{term "1::nat"}
+ − 1191
val ctxt3 = ctxt2 |> update @{term "2::nat"}
328
+ − 1192
in
347
+ − 1193
print ctxt0;
+ − 1194
print ctxt1;
+ − 1195
print ctxt2;
+ − 1196
print ctxt3
569
+ − 1197
end\<close>
+ − 1198
\<open>Empty!
328
+ − 1199
True \<and> True, False
+ − 1200
1
569
+ − 1201
2, 1\<close>}
328
+ − 1202
+ − 1203
Many functions in Isabelle manage and update data in a similar
414
+ − 1204
fashion. Consequently, such calculations with contexts occur frequently in
328
+ − 1205
Isabelle code, although the ``context flow'' is usually only linear.
+ − 1206
Note also that the calculation above has no effect on the underlying
+ − 1207
theory. Once we throw away the contexts, we have no access to their
414
+ − 1208
associated data. This is different for theories, where the command
328
+ − 1209
\isacommand{setup} registers the data with the current and future
330
+ − 1210
theories, and therefore one can access the data potentially
347
+ − 1211
indefinitely.
329
+ − 1212
484
+ − 1213
Move elsewhere
+ − 1214
350
+ − 1215
For convenience there is an abstract layer, namely the type @{ML_type Context.generic},
+ − 1216
for treating theories and proof contexts more uniformly. This type is defined as follows
565
+ − 1217
\<close>
330
+ − 1218
565
+ − 1219
ML_val %grayML\<open>datatype generic =
330
+ − 1220
Theory of theory
565
+ − 1221
| Proof of proof\<close>
330
+ − 1222
565
+ − 1223
text \<open>
350
+ − 1224
\footnote{\bf FIXME: say more about generic contexts.}
329
+ − 1225
+ − 1226
There are two special instances of the data storage mechanism described
350
+ − 1227
above. The first instance implements named theorem lists using the functor
567
+ − 1228
@{ML_functor_ind Named_Thms}. This is because storing theorems in a list
350
+ − 1229
is such a common task. To obtain a named theorem list, you just declare
565
+ − 1230
\<close>
329
+ − 1231
565
+ − 1232
ML %grayML\<open>structure FooRules = Named_Thms
481
+ − 1233
(val name = @{binding "foo"}
565
+ − 1234
val description = "Theorems for foo")\<close>
329
+ − 1235
565
+ − 1236
text \<open>
567
+ − 1237
and set up the @{ML_structure FooRules} with the command
565
+ − 1238
\<close>
329
+ − 1239
565
+ − 1240
setup %gray \<open>FooRules.setup\<close>
329
+ − 1241
565
+ − 1242
text \<open>
329
+ − 1243
This code declares a data container where the theorems are stored,
565
+ − 1244
an attribute \<open>foo\<close> (with the \<open>add\<close> and \<open>del\<close> options
377
+ − 1245
for adding and deleting theorems) and an internal ML-interface for retrieving and
+ − 1246
modifying the theorems.
350
+ − 1247
Furthermore, the theorems are made available on the user-level under the name
565
+ − 1248
\<open>foo\<close>. For example you can declare three lemmas to be a member of the
+ − 1249
theorem list \<open>foo\<close> by:
+ − 1250
\<close>
326
+ − 1251
329
+ − 1252
lemma rule1[foo]: "A" sorry
+ − 1253
lemma rule2[foo]: "B" sorry
+ − 1254
lemma rule3[foo]: "C" sorry
+ − 1255
565
+ − 1256
text \<open>and undeclare the first one by:\<close>
329
+ − 1257
+ − 1258
declare rule1[foo del]
+ − 1259
565
+ − 1260
text \<open>You can query the remaining ones with:
329
+ − 1261
+ − 1262
\begin{isabelle}
565
+ − 1263
\isacommand{thm}~\<open>foo\<close>\\
+ − 1264
\<open>> ?C\<close>\\
+ − 1265
\<open>> ?B\<close>
329
+ − 1266
\end{isabelle}
+ − 1267
+ − 1268
On the ML-level, we can add theorems to the list with @{ML FooRules.add_thm}:
565
+ − 1269
\<close>
329
+ − 1270
565
+ − 1271
setup %gray \<open>Context.theory_map (FooRules.add_thm @{thm TrueI})\<close>
329
+ − 1272
565
+ − 1273
text \<open>
329
+ − 1274
The rules in the list can be retrieved using the function
+ − 1275
@{ML FooRules.get}:
+ − 1276
572
+ − 1277
@{ML_response [display,gray]
569
+ − 1278
\<open>FooRules.get @{context}\<close>
572
+ − 1279
\<open>["True", "?C", "?B"]\<close>}
347
+ − 1280
+ − 1281
Note that this function takes a proof context as argument. This might be
350
+ − 1282
confusing, since the theorem list is stored as theory data. It becomes clear by knowing
+ − 1283
that the proof context contains the information about the current theory and so the function
347
+ − 1284
can access the theorem list in the theory via the context.
329
+ − 1285
+ − 1286
\begin{readmore}
347
+ − 1287
For more information about named theorem lists see
+ − 1288
@{ML_file "Pure/Tools/named_thms.ML"}.
329
+ − 1289
\end{readmore}
+ − 1290
+ − 1291
The second special instance of the data storage mechanism are configuration
+ − 1292
values. They are used to enable users to configure tools without having to
+ − 1293
resort to the ML-level (and also to avoid references). Assume you want the
565
+ − 1294
user to control three values, say \<open>bval\<close> containing a boolean, \<open>ival\<close> containing an integer and \<open>sval\<close> containing a string. These
329
+ − 1295
values can be declared by
565
+ − 1296
\<close>
329
+ − 1297
565
+ − 1298
ML %grayML\<open>val bval = Attrib.setup_config_bool @{binding "bval"} (K false)
462
+ − 1299
val ival = Attrib.setup_config_int @{binding "ival"} (K 0)
565
+ − 1300
val sval = Attrib.setup_config_string @{binding "sval"} (K "some string")\<close>
329
+ − 1301
565
+ − 1302
text \<open>
462
+ − 1303
where each value needs to be given a default.
329
+ − 1304
The user can now manipulate the values from the user-level of Isabelle
+ − 1305
with the command
565
+ − 1306
\<close>
329
+ − 1307
+ − 1308
declare [[bval = true, ival = 3]]
+ − 1309
565
+ − 1310
text \<open>
329
+ − 1311
On the ML-level these values can be retrieved using the
346
+ − 1312
function @{ML_ind get in Config} from a proof context
+ − 1313
567
+ − 1314
@{ML_matchresult [display,gray]
569
+ − 1315
\<open>Config.get @{context} bval\<close>
+ − 1316
\<open>true\<close>}
346
+ − 1317
423
+ − 1318
or directly from a theory using the function @{ML_ind get_global in Config}
346
+ − 1319
567
+ − 1320
@{ML_matchresult [display,gray]
569
+ − 1321
\<open>Config.get_global @{theory} bval\<close>
+ − 1322
\<open>true\<close>}
329
+ − 1323
346
+ − 1324
It is also possible to manipulate the configuration values
347
+ − 1325
from the ML-level with the functions @{ML_ind put in Config}
423
+ − 1326
and @{ML_ind put_global in Config}. For example
346
+ − 1327
567
+ − 1328
@{ML_matchresult [display,gray]
569
+ − 1329
\<open>let
346
+ − 1330
val ctxt = @{context}
569
+ − 1331
val ctxt' = Config.put sval "foo" ctxt
+ − 1332
val ctxt'' = Config.put sval "bar" ctxt'
346
+ − 1333
in
350
+ − 1334
(Config.get ctxt sval,
+ − 1335
Config.get ctxt' sval,
+ − 1336
Config.get ctxt'' sval)
569
+ − 1337
end\<close>
+ − 1338
\<open>("some string", "foo", "bar")\<close>}
329
+ − 1339
468
+ − 1340
A concrete example for a configuration value is
+ − 1341
@{ML_ind simp_trace in Raw_Simplifier}, which switches on trace information
474
+ − 1342
in the simplifier. This can be used for example in the following proof
565
+ − 1343
\<close>
468
+ − 1344
562
+ − 1345
468
+ − 1346
lemma
+ − 1347
shows "(False \<or> True) \<and> True"
+ − 1348
proof (rule conjI)
+ − 1349
show "False \<or> True" using [[simp_trace = true]] by simp
+ − 1350
next
+ − 1351
show "True" by simp
+ − 1352
qed
+ − 1353
565
+ − 1354
text \<open>
468
+ − 1355
in order to inspect how the simplifier solves the first subgoal.
+ − 1356
329
+ − 1357
\begin{readmore}
+ − 1358
For more information about configuration values see
346
+ − 1359
the files @{ML_file "Pure/Isar/attrib.ML"} and
+ − 1360
@{ML_file "Pure/config.ML"}.
329
+ − 1361
\end{readmore}
565
+ − 1362
\<close>
343
+ − 1363
565
+ − 1364
section \<open>Summary\<close>
343
+ − 1365
565
+ − 1366
text \<open>
343
+ − 1367
This chapter describes the combinators that are used in Isabelle, as well
+ − 1368
as a simple printing infrastructure for @{ML_type term}, @{ML_type cterm}
+ − 1369
and @{ML_type thm}. The section on ML-antiquotations shows how to refer
+ − 1370
statically to entities from the logic level of Isabelle. Isabelle also
+ − 1371
contains mechanisms for storing arbitrary data in theory and proof
+ − 1372
contexts.
+ − 1373
347
+ − 1374
\begin{conventions}
+ − 1375
\begin{itemize}
370
+ − 1376
\item Print messages that belong together in a single string.
350
+ − 1377
\item Do not use references in Isabelle code.
347
+ − 1378
\end{itemize}
+ − 1379
\end{conventions}
+ − 1380
565
+ − 1381
\<close>
196
+ − 1382
end