2
+ − 1
theory FirstSteps
25
+ − 2
imports Base
2
+ − 3
begin
+ − 4
+ − 5
chapter {* First Steps *}
+ − 6
255
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 7
42
+ − 8
text {*
242
+ − 9
Isabelle programming is done in ML. Just like lemmas and proofs, ML-code
89
+ − 10
in Isabelle is part of a theory. If you want to follow the code given in
54
+ − 11
this chapter, we assume you are working inside the theory starting with
2
+ − 12
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 13
\begin{quote}
5
+ − 14
\begin{tabular}{@ {}l}
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 15
\isacommand{theory} FirstSteps\\
5
+ − 16
\isacommand{imports} Main\\
+ − 17
\isacommand{begin}\\
6
+ − 18
\ldots
5
+ − 19
\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
+ − 20
\end{quote}
157
+ − 21
+ − 22
We also generally assume you are working with HOL. The given examples might
193
+ − 23
need to be adapted if you work in a different logic.
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 24
*}
5
+ − 25
20
+ − 26
section {* Including ML-Code *}
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 27
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 28
text {*
5
+ − 29
The easiest and quickest way to include code in a theory is
89
+ − 30
by using the \isacommand{ML}-command. For example:
2
+ − 31
75
+ − 32
\begin{isabelle}
+ − 33
\begin{graybox}
85
+ − 34
\isacommand{ML}~@{text "\<verbopen>"}\isanewline
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 35
\hspace{5mm}@{ML "3 + 4"}\isanewline
85
+ − 36
@{text "\<verbclose>"}\isanewline
+ − 37
@{text "> 7"}\smallskip
75
+ − 38
\end{graybox}
+ − 39
\end{isabelle}
2
+ − 40
234
+ − 41
Like normal Isabelle scripts, \isacommand{ML}-commands can be
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
+ − 42
evaluated by using the advance and undo buttons of your Isabelle
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 43
environment. The code inside the \isacommand{ML}-command can also contain
193
+ − 44
value and function bindings, for example
+ − 45
*}
+ − 46
+ − 47
ML %gray {*
+ − 48
val r = ref 0
+ − 49
fun f n = n + 1
+ − 50
*}
+ − 51
+ − 52
text {*
+ − 53
and even those can be undone when the proof
+ − 54
script is retracted. As mentioned in the Introduction, we will drop the
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
+ − 55
\isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"} scaffolding whenever we
107
+ − 56
show code. The lines prefixed with @{text [quotes] ">"} are not part of the
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
+ − 57
code, rather they indicate what the response is when the code is evaluated.
254
+ − 58
There are also the commands \isacommand{ML\_val} and \isacommand{ML\_prf} for
+ − 59
including ML-code. The first evaluates the given code, but any effect on the
264
+ − 60
theory, in which the code is embedded, is suppressed. The second needs to
+ − 61
be used if ML-code is defined
253
+ − 62
inside a proof. For example
+ − 63
254
+ − 64
\begin{quote}
+ − 65
\begin{isabelle}
+ − 66
\isacommand{lemma}~@{text "test:"}\isanewline
+ − 67
\isacommand{shows}~@{text [quotes] "True"}\isanewline
+ − 68
\isacommand{ML\_prf}~@{text "\<verbopen>"}~@{ML "writeln \"Trivial!\""}~@{text "\<verbclose>"}\isanewline
+ − 69
\isacommand{oops}
+ − 70
\end{isabelle}
+ − 71
\end{quote}
253
+ − 72
263
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 73
However, both commands will only play minor roles in this tutorial (we will always
264
+ − 74
arrange that the ML-code is defined outside of proofs, for example).
10
+ − 75
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
+ − 76
Once a portion of code is relatively stable, you usually want to export it
235
+ − 77
to a separate ML-file. Such files can then be included somewhere inside a
+ − 78
theory by using the command \isacommand{use}. For example
+ − 79
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 80
\begin{quote}
235
+ − 81
\begin{tabular}{@ {}l}
+ − 82
\isacommand{theory} FirstSteps\\
+ − 83
\isacommand{imports} Main\\
+ − 84
\isacommand{uses}~@{text "(\"file_to_be_included.ML\")"} @{text "\<dots>"}\\
+ − 85
\isacommand{begin}\\
+ − 86
\ldots\\
+ − 87
\isacommand{use}~@{text "\"file_to_be_included.ML\""}\\
+ − 88
\ldots
+ − 89
\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
+ − 90
\end{quote}
235
+ − 91
+ − 92
The \isacommand{uses}-command in the header of the theory is needed in order
+ − 93
to indicate the dependency of the theory on the ML-file. Alternatively, the
+ − 94
file can be included by just writing in the header
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 95
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 96
\begin{quote}
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 97
\begin{tabular}{@ {}l}
54
+ − 98
\isacommand{theory} FirstSteps\\
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 99
\isacommand{imports} Main\\
58
+ − 100
\isacommand{uses} @{text "\"file_to_be_included.ML\""} @{text "\<dots>"}\\
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 101
\isacommand{begin}\\
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 102
\ldots
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 103
\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
+ − 104
\end{quote}
235
+ − 105
263
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 106
Note that no parentheses are given this time. Note also that the
264
+ − 107
included ML-file should not contain any
263
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 108
\isacommand{use} itself. Otherwise Isabelle is unable to record all
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 109
file dependencies, which is a nuisance if you have to track down
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 110
errors.
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 111
*}
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 112
126
+ − 113
section {* Debugging and Printing\label{sec:printing} *}
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 114
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 115
text {*
13
+ − 116
50
+ − 117
During development you might find it necessary to inspect some data
10
+ − 118
in your code. This can be done in a ``quick-and-dirty'' fashion using
256
+ − 119
the function @{ML [index] "writeln"}. For example
10
+ − 120
239
+ − 121
@{ML_response_fake [display,gray] "writeln \"any string\"" "\"any string\""}
10
+ − 122
58
+ − 123
will print out @{text [quotes] "any string"} inside the response buffer
+ − 124
of Isabelle. This function expects a string as argument. If you develop under PolyML,
50
+ − 125
then there is a convenient, though again ``quick-and-dirty'', method for
240
+ − 126
converting values into strings, namely the function @{ML PolyML.makestring}:
10
+ − 127
240
+ − 128
@{ML_response_fake [display,gray] "writeln (PolyML.makestring 1)" "\"1\""}
12
+ − 129
260
+ − 130
However, @{ML [index] makestring in PolyML} only works if the type of what is converted is monomorphic
78
+ − 131
and not a function.
12
+ − 132
239
+ − 133
The function @{ML "writeln"} should only be used for testing purposes, because any
52
+ − 134
output this function generates will be overwritten as soon as an error is
50
+ − 135
raised. For printing anything more serious and elaborate, the
256
+ − 136
function @{ML [index] tracing} is more appropriate. This function writes all output into
89
+ − 137
a separate tracing buffer. For example:
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 138
161
+ − 139
@{ML_response_fake [display,gray] "tracing \"foo\"" "\"foo\""}
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 140
58
+ − 141
It is also possible to redirect the ``channel'' where the string @{text "foo"} is
196
+ − 142
printed to a separate file, e.g., to prevent ProofGeneral from choking on massive
107
+ − 143
amounts of trace output. This redirection can be achieved with the code:
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 144
*}
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 145
69
+ − 146
ML{*val strip_specials =
42
+ − 147
let
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 148
fun strip ("\^A" :: _ :: cs) = strip cs
42
+ − 149
| strip (c :: cs) = c :: strip cs
+ − 150
| strip [] = [];
+ − 151
in implode o strip o explode end;
+ − 152
+ − 153
fun redirect_tracing stream =
+ − 154
Output.tracing_fn := (fn s =>
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 155
(TextIO.output (stream, (strip_specials s));
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 156
TextIO.output (stream, "\n");
69
+ − 157
TextIO.flushOut stream)) *}
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 158
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 159
text {*
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 160
Calling @{ML "redirect_tracing"} with @{ML "(TextIO.openOut \"foo.bar\")"}
58
+ − 161
will cause that all tracing information is printed into the file @{text "foo.bar"}.
75
+ − 162
256
+ − 163
You can print out error messages with the function @{ML [index] error}; for example:
75
+ − 164
122
+ − 165
@{ML_response_fake [display,gray] "if 0=1 then true else (error \"foo\")"
+ − 166
"Exception- ERROR \"foo\" raised
+ − 167
At command \"ML\"."}
75
+ − 168
256
+ − 169
(FIXME Mention how to work with @{ML [index] debug in Toplevel} and
+ − 170
@{ML [index] profiling in Toplevel}.)
192
+ − 171
*}
+ − 172
+ − 173
(*
207
+ − 174
ML {* reset Toplevel.debug *}
+ − 175
+ − 176
ML {* fun dodgy_fun () = (raise TYPE ("",[],[]); 1) *}
193
+ − 177
207
+ − 178
ML {* fun innocent () = dodgy_fun () *}
+ − 179
ML {* exception_trace (fn () => cterm_of @{theory} (Bound 0)) *}
+ − 180
ML {* exception_trace (fn () => innocent ()) *}
192
+ − 181
207
+ − 182
ML {* Toplevel.program (fn () => cterm_of @{theory} (Bound 0)) *}
192
+ − 183
207
+ − 184
ML {* Toplevel.program (fn () => innocent ()) *}
192
+ − 185
*)
+ − 186
+ − 187
text {*
126
+ − 188
Most often you want to inspect data of type @{ML_type term}, @{ML_type cterm}
251
+ − 189
or @{ML_type thm}. Isabelle contains elaborate pretty-printing functions for printing
+ − 190
them (see Section \ref{sec:pretty}),
126
+ − 191
but for quick-and-dirty solutions they are far too unwieldy. A simple way to transform
256
+ − 192
a term into a string is to use the function @{ML [index] string_of_term in Syntax}.
126
+ − 193
+ − 194
@{ML_response_fake [display,gray]
+ − 195
"Syntax.string_of_term @{context} @{term \"1::nat\"}"
+ − 196
"\"\\^E\\^Fterm\\^E\\^E\\^Fconst\\^Fname=HOL.one_class.one\\^E1\\^E\\^F\\^E\\^E\\^F\\^E\""}
+ − 197
+ − 198
This produces a string with some additional information encoded in it. The string
256
+ − 199
can be properly printed by using the function @{ML [index] writeln}.
126
+ − 200
+ − 201
@{ML_response_fake [display,gray]
239
+ − 202
"writeln (Syntax.string_of_term @{context} @{term \"1::nat\"})"
126
+ − 203
"\"1\""}
+ − 204
+ − 205
A @{ML_type cterm} can be transformed into a string by the following function.
+ − 206
*}
+ − 207
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 208
ML{*fun string_of_cterm ctxt t =
126
+ − 209
Syntax.string_of_term ctxt (term_of t)*}
+ − 210
+ − 211
text {*
256
+ − 212
In this example the function @{ML [index] term_of} extracts the @{ML_type term} from
149
+ − 213
a @{ML_type cterm}. If there are more than one @{ML_type cterm}s to be
256
+ − 214
printed, you can use the function @{ML [index] commas} to separate them.
126
+ − 215
*}
+ − 216
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 217
ML{*fun string_of_cterms ctxt ts =
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 218
commas (map (string_of_cterm ctxt) ts)*}
126
+ − 219
+ − 220
text {*
+ − 221
The easiest way to get the string of a theorem is to transform it
256
+ − 222
into a @{ML_type cterm} using the function @{ML [index] crep_thm}.
190
+ − 223
*}
+ − 224
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 225
ML{*fun string_of_thm ctxt thm =
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 226
string_of_cterm ctxt (#prop (crep_thm thm))*}
190
+ − 227
+ − 228
text {*
+ − 229
Theorems also include schematic variables, such as @{text "?P"},
+ − 230
@{text "?Q"} and so on.
+ − 231
+ − 232
@{ML_response_fake [display, gray]
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 233
"writeln (string_of_thm @{context} @{thm conjI})"
190
+ − 234
"\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q"}
+ − 235
+ − 236
In order to improve the readability of theorems we convert
158
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 237
these schematic variables into free variables using the
263
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 238
function @{ML [index] import in Variable}.
126
+ − 239
*}
+ − 240
158
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 241
ML{*fun no_vars ctxt thm =
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 242
let
263
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 243
val ((_, [thm']), _) = Variable.import true [thm] ctxt
158
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 244
in
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 245
thm'
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 246
end
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 247
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 248
fun string_of_thm_no_vars ctxt thm =
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 249
string_of_cterm ctxt (#prop (crep_thm (no_vars ctxt thm)))*}
126
+ − 250
+ − 251
text {*
207
+ − 252
Theorem @{thm [source] conjI} is now printed as follows:
190
+ − 253
+ − 254
@{ML_response_fake [display, gray]
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 255
"writeln (string_of_thm_no_vars @{context} @{thm conjI})"
196
+ − 256
"\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> P \<and> Q"}
190
+ − 257
126
+ − 258
Again the function @{ML commas} helps with printing more than one theorem.
+ − 259
*}
+ − 260
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 261
ML{*fun string_of_thms ctxt thms =
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 262
commas (map (string_of_thm ctxt) thms)
190
+ − 263
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 264
fun string_of_thms_no_vars ctxt thms =
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 265
commas (map (string_of_thm_no_vars ctxt) thms) *}
126
+ − 266
+ − 267
section {* Combinators\label{sec:combinators} *}
+ − 268
+ − 269
text {*
131
+ − 270
For beginners perhaps the most puzzling parts in the existing code of Isabelle are
126
+ − 271
the combinators. At first they seem to greatly obstruct the
+ − 272
comprehension of the code, but after getting familiar with them, they
+ − 273
actually ease the understanding and also the programming.
+ − 274
256
+ − 275
The simplest combinator is @{ML [index] I}, which is just the identity function defined as
126
+ − 276
*}
+ − 277
+ − 278
ML{*fun I x = x*}
+ − 279
256
+ − 280
text {* Another simple combinator is @{ML [index] K}, defined as *}
126
+ − 281
+ − 282
ML{*fun K x = fn _ => x*}
+ − 283
+ − 284
text {*
256
+ − 285
@{ML [index] K} ``wraps'' a function around the argument @{text "x"}. However, this
126
+ − 286
function ignores its argument. As a result, @{ML K} defines a constant function
+ − 287
always returning @{text x}.
+ − 288
256
+ − 289
The next combinator is reverse application, @{ML [index] "|>"}, defined as:
126
+ − 290
*}
+ − 291
+ − 292
ML{*fun x |> f = f x*}
+ − 293
+ − 294
text {* While just syntactic sugar for the usual function application,
+ − 295
the purpose of this combinator is to implement functions in a
+ − 296
``waterfall fashion''. Consider for example the function *}
+ − 297
+ − 298
ML %linenosgray{*fun inc_by_five x =
+ − 299
x |> (fn x => x + 1)
+ − 300
|> (fn x => (x, x))
+ − 301
|> fst
+ − 302
|> (fn x => x + 4)*}
+ − 303
+ − 304
text {*
196
+ − 305
which increments its argument @{text x} by 5. It proceeds by first incrementing
126
+ − 306
the argument by 1 (Line 2); then storing the result in a pair (Line 3); taking
+ − 307
the first component of the pair (Line 4) and finally incrementing the first
+ − 308
component by 4 (Line 5). This kind of cascading manipulations of values is quite
+ − 309
common when dealing with theories (for example by adding a definition, followed by
+ − 310
lemmas and so on). The reverse application allows you to read what happens in
+ − 311
a top-down manner. This kind of coding should also be familiar,
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 312
if you have been exposed to Haskell's {\it do}-notation. Writing the function
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 313
@{ML inc_by_five} using the reverse application is much clearer than writing
126
+ − 314
*}
+ − 315
+ − 316
ML{*fun inc_by_five x = fst ((fn x => (x, x)) (x + 1)) + 4*}
+ − 317
+ − 318
text {* or *}
+ − 319
+ − 320
ML{*fun inc_by_five x =
+ − 321
((fn x => x + 4) o fst o (fn x => (x, x)) o (fn x => x + 1)) x*}
+ − 322
+ − 323
text {* and typographically more economical than *}
+ − 324
+ − 325
ML{*fun inc_by_five x =
257
+ − 326
let val y1 = x + 1
+ − 327
val y2 = (y1, y1)
+ − 328
val y3 = fst y2
+ − 329
val y4 = y3 + 4
+ − 330
in y4 end*}
126
+ − 331
+ − 332
text {*
+ − 333
Another reason why the let-bindings in the code above are better to be
+ − 334
avoided: it is more than easy to get the intermediate values wrong, not to
+ − 335
mention the nightmares the maintenance of this code causes!
+ − 336
215
+ − 337
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
+ − 338
the waterfall fashion might be the following code:
177
+ − 339
*}
126
+ − 340
193
+ − 341
ML %linenosgray{*fun apply_fresh_args f ctxt =
+ − 342
f |> fastype_of
+ − 343
|> binder_types
+ − 344
|> map (pair "z")
+ − 345
|> Variable.variant_frees ctxt [f]
+ − 346
|> map Free
257
+ − 347
|> curry list_comb f *}
126
+ − 348
177
+ − 349
text {*
266
+ − 350
This function takes a term and a context as argument. If the term is of function
+ − 351
type, then @{ML "apply_fresh_args"} returns the term with distinct variables
298
+ − 352
applied to it. For example below variables are applied to the term
+ − 353
@{term [show_types] "P::nat \<Rightarrow> int \<Rightarrow> unit \<Rightarrow> bool"}:
183
+ − 354
+ − 355
@{ML_response_fake [display,gray]
266
+ − 356
"let
+ − 357
val t = @{term \"P::nat \<Rightarrow> int \<Rightarrow> unit \<Rightarrow> bool\"}
+ − 358
val ctxt = @{context}
+ − 359
in
+ − 360
apply_fresh_args t ctxt
+ − 361
|> Syntax.string_of_term ctxt
+ − 362
|> writeln
+ − 363
end"
183
+ − 364
"P z za zb"}
177
+ − 365
184
+ − 366
You can read off this behaviour from how @{ML apply_fresh_args} is
256
+ − 367
coded: in Line 2, the function @{ML [index] fastype_of} calculates the type of the
298
+ − 368
term; @{ML [index] binder_types} in the next line produces the list of argument
184
+ − 369
types (in the case above the list @{text "[nat, int, unit]"}); Line 4
+ − 370
pairs up each type with the string @{text "z"}; the
256
+ − 371
function @{ML [index] variant_frees in Variable} generates for each @{text "z"} a
193
+ − 372
unique name avoiding the given @{text f}; the list of name-type pairs is turned
184
+ − 373
into a list of variable terms in Line 6, which in the last line is applied
298
+ − 374
by the function @{ML [index] list_comb} to the term. In this last step we have to
256
+ − 375
use the function @{ML [index] curry}, because @{ML [index] list_comb} expects the
+ − 376
function and the variables list as a pair. This kind of functions is often needed when
298
+ − 377
constructing terms with fresh variables. The infrastructure helps tremendously
+ − 378
to avoid any name clashes. Consider for example:
252
+ − 379
+ − 380
@{ML_response_fake [display,gray]
266
+ − 381
"let
+ − 382
val t = @{term \"za::'a \<Rightarrow> 'b \<Rightarrow> 'c\"}
+ − 383
val ctxt = @{context}
+ − 384
in
+ − 385
apply_fresh_args t ctxt
+ − 386
|> Syntax.string_of_term ctxt
+ − 387
|> writeln
+ − 388
end"
252
+ − 389
"za z zb"}
177
+ − 390
266
+ − 391
where the @{text "za"} is correctly avoided.
+ − 392
256
+ − 393
The combinator @{ML [index] "#>"} is the reverse function composition. It can be
177
+ − 394
used to define the following function
126
+ − 395
*}
+ − 396
+ − 397
ML{*val inc_by_six =
+ − 398
(fn x => x + 1)
+ − 399
#> (fn x => x + 2)
+ − 400
#> (fn x => x + 3)*}
+ − 401
+ − 402
text {*
+ − 403
which is the function composed of first the increment-by-one function and then
+ − 404
increment-by-two, followed by increment-by-three. Again, the reverse function
+ − 405
composition allows you to read the code top-down.
+ − 406
+ − 407
The remaining combinators described in this section add convenience for the
256
+ − 408
``waterfall method'' of writing functions. The combinator @{ML [index] tap} allows
126
+ − 409
you to get hold of an intermediate result (to do some side-calculations for
+ − 410
instance). The function
+ − 411
+ − 412
*}
+ − 413
+ − 414
ML %linenosgray{*fun inc_by_three x =
+ − 415
x |> (fn x => x + 1)
240
+ − 416
|> tap (fn x => tracing (PolyML.makestring x))
126
+ − 417
|> (fn x => x + 2)*}
+ − 418
+ − 419
text {*
+ − 420
increments the argument first by @{text "1"} and then by @{text "2"}. In the
256
+ − 421
middle (Line 3), however, it uses @{ML [index] tap} for printing the ``plus-one''
+ − 422
intermediate result inside the tracing buffer. The function @{ML [index] tap} can
126
+ − 423
only be used for side-calculations, because any value that is computed
+ − 424
cannot be merged back into the ``main waterfall''. To do this, you can use
+ − 425
the next combinator.
+ − 426
256
+ − 427
The combinator @{ML [index] "`"} (a backtick) is similar to @{ML tap}, but applies a
196
+ − 428
function to the value and returns the result together with the value (as a
+ − 429
pair). For example the function
126
+ − 430
*}
+ − 431
+ − 432
ML{*fun inc_as_pair x =
+ − 433
x |> `(fn x => x + 1)
+ − 434
|> (fn (x, y) => (x, y + 1))*}
+ − 435
+ − 436
text {*
+ − 437
takes @{text x} as argument, and then increments @{text x}, but also keeps
+ − 438
@{text x}. The intermediate result is therefore the pair @{ML "(x + 1, x)"
+ − 439
for x}. After that, the function increments the right-hand component of the
+ − 440
pair. So finally the result will be @{ML "(x + 1, x + 1)" for x}.
+ − 441
256
+ − 442
The combinators @{ML [index] "|>>"} and @{ML [index] "||>"} are defined for
126
+ − 443
functions manipulating pairs. The first applies the function to
+ − 444
the first component of the pair, defined as
+ − 445
*}
+ − 446
+ − 447
ML{*fun (x, y) |>> f = (f x, y)*}
+ − 448
+ − 449
text {*
+ − 450
and the second combinator to the second component, defined as
+ − 451
*}
+ − 452
+ − 453
ML{*fun (x, y) ||> f = (x, f y)*}
+ − 454
+ − 455
text {*
256
+ − 456
With the combinator @{ML [index] "|->"} you can re-combine the elements from a pair.
126
+ − 457
This combinator is defined as
+ − 458
*}
+ − 459
+ − 460
ML{*fun (x, y) |-> f = f x y*}
+ − 461
215
+ − 462
text {*
+ − 463
and can be used to write the following roundabout version
126
+ − 464
of the @{text double} function:
+ − 465
*}
+ − 466
+ − 467
ML{*fun double x =
+ − 468
x |> (fn x => (x, x))
+ − 469
|-> (fn x => fn y => x + y)*}
+ − 470
215
+ − 471
text {*
256
+ − 472
The combinator @{ML [index] ||>>} plays a central rôle whenever your task is to update a
215
+ − 473
theory and the update also produces a side-result (for example a theorem). Functions
+ − 474
for such tasks return a pair whose second component is the theory and the fist
256
+ − 475
component is the side-result. Using @{ML [index] ||>>}, you can do conveniently the update
280
+ − 476
and also accumulate the side-results. Consider the following simple function.
215
+ − 477
*}
+ − 478
+ − 479
ML %linenosgray{*fun acc_incs x =
+ − 480
x |> (fn x => ("", x))
+ − 481
||>> (fn x => (x, x + 1))
+ − 482
||>> (fn x => (x, x + 1))
+ − 483
||>> (fn x => (x, x + 1))*}
+ − 484
+ − 485
text {*
+ − 486
The purpose of Line 2 is to just pair up the argument with a dummy value (since
256
+ − 487
@{ML [index] "||>>"} operates on pairs). Each of the next three lines just increment
280
+ − 488
the value by one, but also nest the intermediate results to the left. For example
215
+ − 489
+ − 490
@{ML_response [display,gray]
+ − 491
"acc_incs 1"
+ − 492
"((((\"\", 1), 2), 3), 4)"}
+ − 493
+ − 494
You can continue this chain with:
+ − 495
+ − 496
@{ML_response [display,gray]
+ − 497
"acc_incs 1 ||>> (fn x => (x, x + 2))"
+ − 498
"(((((\"\", 1), 2), 3), 4), 6)"}
+ − 499
298
+ − 500
(FIXME: maybe give a ``real world'' example for this combinator)
215
+ − 501
*}
+ − 502
126
+ − 503
text {*
256
+ − 504
Recall that @{ML [index] "|>"} is the reverse function application. Recall also that
196
+ − 505
the related
256
+ − 506
reverse function composition is @{ML [index] "#>"}. In fact all the combinators
+ − 507
@{ML [index] "|->"}, @{ML [index] "|>>"} , @{ML [index] "||>"} and @{ML [index]
+ − 508
"||>>"} described above have related combinators for
+ − 509
function composition, namely @{ML [index] "#->"}, @{ML [index] "#>>"},
+ − 510
@{ML [index] "##>"} and @{ML [index] "##>>"}.
+ − 511
Using @{ML [index] "#->"}, for example, the function @{text double} can also be written as:
126
+ − 512
*}
+ − 513
+ − 514
ML{*val double =
+ − 515
(fn x => (x, x))
+ − 516
#-> (fn x => fn y => x + y)*}
+ − 517
+ − 518
text {*
+ − 519
+ − 520
(FIXME: find a good exercise for combinators)
127
+ − 521
+ − 522
\begin{readmore}
196
+ − 523
The most frequently used combinators are defined in the files @{ML_file
+ − 524
"Pure/library.ML"}
127
+ − 525
and @{ML_file "Pure/General/basics.ML"}. Also \isccite{sec:ML-linear-trans}
+ − 526
contains further information about combinators.
+ − 527
\end{readmore}
126
+ − 528
251
+ − 529
(FIXME: say something about calling conventions)
15
+ − 530
*}
+ − 531
10
+ − 532
2
+ − 533
section {* Antiquotations *}
+ − 534
+ − 535
text {*
49
+ − 536
The main advantage of embedding all code in a theory is that the code can
58
+ − 537
contain references to entities defined on the logical level of Isabelle. By
+ − 538
this we mean definitions, theorems, terms and so on. This kind of reference is
298
+ − 539
realised with antiquotations, sometimes also referred to as ML-antiquotations.
+ − 540
For example, one can print out the name of the current
49
+ − 541
theory by typing
+ − 542
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 543
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 544
@{ML_response [display,gray] "Context.theory_name @{theory}" "\"FirstSteps\""}
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 545
5
+ − 546
where @{text "@{theory}"} is an antiquotation that is substituted with the
49
+ − 547
current theory (remember that we assumed we are inside the theory
89
+ − 548
@{text FirstSteps}). The name of this theory can be extracted using
256
+ − 549
the function @{ML [index] theory_name in Context}.
5
+ − 550
89
+ − 551
Note, however, that antiquotations are statically linked, that is their value is
12
+ − 552
determined at ``compile-time'', not ``run-time''. For example the function
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 553
*}
5
+ − 554
69
+ − 555
ML{*fun not_current_thyname () = Context.theory_name @{theory} *}
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 556
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 557
text {*
2
+ − 558
89
+ − 559
does \emph{not} return the name of the current theory, if it is run in a
5
+ − 560
different theory. Instead, the code above defines the constant function
58
+ − 561
that always returns the string @{text [quotes] "FirstSteps"}, no matter where the
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 562
function is called. Operationally speaking, the antiquotation @{text "@{theory}"} is
5
+ − 563
\emph{not} replaced with code that will look up the current theory in
+ − 564
some data structure and return it. Instead, it is literally
+ − 565
replaced with the value representing the theory name.
2
+ − 566
132
+ − 567
In a similar way you can use antiquotations to refer to proved theorems:
133
+ − 568
@{text "@{thm \<dots>}"} for a single theorem
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 569
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 570
@{ML_response_fake [display,gray] "@{thm allI}" "(\<And>x. ?P x) \<Longrightarrow> \<forall>x. ?P x"}
75
+ − 571
133
+ − 572
and @{text "@{thms \<dots>}"} for more than one
132
+ − 573
+ − 574
@{ML_response_fake [display,gray] "@{thms conj_ac}"
+ − 575
"(?P \<and> ?Q) = (?Q \<and> ?P)
+ − 576
(?P \<and> ?Q \<and> ?R) = (?Q \<and> ?P \<and> ?R)
+ − 577
((?P \<and> ?Q) \<and> ?R) = (?P \<and> ?Q \<and> ?R)"}
+ − 578
292
+ − 579
The point of these antiquotations is that referring to theorems in this way
+ − 580
makes your code independent from what theorems the user might have stored
+ − 581
under this name (this becomes especially important when you deal with
+ − 582
theorem lists; see Section \ref{sec:attributes}).
+ − 583
+ − 584
You can also refer to the current simpset via an antiquotation. To illustrate
+ − 585
this we implement the function that extracts the theorem names stored in a
+ − 586
simpset.
131
+ − 587
*}
75
+ − 588
149
+ − 589
ML{*fun get_thm_names_from_ss simpset =
131
+ − 590
let
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
+ − 591
val {simps,...} = MetaSimplifier.dest_ss simpset
70
+ − 592
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
+ − 593
map #1 simps
131
+ − 594
end*}
54
+ − 595
131
+ − 596
text {*
256
+ − 597
The function @{ML [index] dest_ss in MetaSimplifier} returns a record containing all
193
+ − 598
information stored in the simpset, but 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
+ − 599
simp-rules. Now you can feed in the current simpset into this function.
193
+ − 600
The current simpset can be referred to using the antiquotation @{text "@{simpset}"}.
81
+ − 601
131
+ − 602
@{ML_response_fake [display,gray]
149
+ − 603
"get_thm_names_from_ss @{simpset}"
+ − 604
"[\"Nat.of_nat_eq_id\", \"Int.of_int_eq_id\", \"Nat.One_nat_def\", \<dots>]"}
10
+ − 605
196
+ − 606
Again, this way of referencing simpsets makes you independent from additions
292
+ − 607
of lemmas to the simpset by the user that can potentially cause loops in your
+ − 608
code.
156
+ − 609
251
+ − 610
On the ML-level of Isabelle, you often have to work with qualified names.
+ − 611
These are strings with some additional information, such as positional information
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 612
and qualifiers. Such qualified names can be generated with the antiquotation
292
+ − 613
@{text "@{binding \<dots>}"}. For example
192
+ − 614
+ − 615
@{ML_response [display,gray]
+ − 616
"@{binding \"name\"}"
+ − 617
"name"}
+ − 618
292
+ − 619
An example where a qualified name is needed is the function
+ − 620
@{ML [index] define in LocalTheory}. This function is used below to define
+ − 621
the constant @{term "TrueConj"} as the conjunction @{term "True \<and> True"}.
192
+ − 622
*}
+ − 623
+ − 624
local_setup %gray {*
+ − 625
snd o LocalTheory.define Thm.internalK
+ − 626
((@{binding "TrueConj"}, NoSyn),
+ − 627
(Attrib.empty_binding, @{term "True \<and> True"})) *}
+ − 628
193
+ − 629
text {*
+ − 630
Now querying the definition you obtain:
+ − 631
+ − 632
\begin{isabelle}
+ − 633
\isacommand{thm}~@{text "TrueConj_def"}\\
225
+ − 634
@{text "> "}~@{thm TrueConj_def}
193
+ − 635
\end{isabelle}
+ − 636
194
+ − 637
(FIXME give a better example why bindings are important; maybe
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 638
give a pointer to \isacommand{local\_setup}; if not, then explain
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 639
why @{ML snd} is needed)
193
+ − 640
292
+ − 641
It is also possible to define your own antiquotations. But you should
298
+ − 642
exercise care when introducing new ones, as they can also make your
+ − 643
code also difficult to read. In the next section we will see how the (build in)
292
+ − 644
antiquotation @{text "@{term \<dots>}"} can be used to construct terms.
+ − 645
A restriction of this antiquotation is that it does not allow you to
298
+ − 646
use schematic variables. If you want an antiquotation that does not
292
+ − 647
have this restriction, you can implement your own using the
298
+ − 648
function @{ML [index] ML_Antiquote.inline}. The code is as follows.
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 649
*}
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 650
292
+ − 651
ML %linenosgray{*ML_Antiquote.inline "term_pat"
263
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 652
(Args.context -- Scan.lift Args.name_source >>
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 653
(fn (ctxt, s) =>
264
+ − 654
s |> ProofContext.read_term_pattern ctxt
+ − 655
|> ML_Syntax.print_term
+ − 656
|> ML_Syntax.atomic))*}
263
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 657
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 658
text {*
298
+ − 659
We call the new antiquotation @{text "term_pat"} (Line 1); the parser in Line
292
+ − 660
2 provides us with a context and a string; this string is transformed into a
298
+ − 661
term using the function @{ML [index] read_term_pattern in ProofContext} (Line 4);
292
+ − 662
the next two lines print the term so that the ML-system can understand
+ − 663
them. An example of this antiquotation is as follows.
+ − 664
+ − 665
@{ML_response_fake [display,gray]
+ − 666
"@{term_pat \"Suc (?x::nat)\"}"
+ − 667
"Const (\"Suc\", \"nat \<Rightarrow> nat\") $ Var ((\"x\", 0), \"nat\")"}
+ − 668
298
+ − 669
which is the internal representation of the term @{text "Suc ?x"}.
+ − 670
263
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 671
\begin{readmore}
292
+ − 672
The file @{ML_file "Pure/ML/ml_antiquote.ML"} contains the the definitions
+ − 673
for most antiquotations.
263
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 674
\end{readmore}
292
+ − 675
298
+ − 676
Note one source of possible confusion about antiquotations. There are two kinds
+ − 677
of them in Isabelle, which have very different purpose and infrastructures. The
+ − 678
first kind, described in this section, are \emph{ML-antiquotations}. They are
+ − 679
used to refer to entities (like terms, types etc) from Isabelle's logic layer
+ − 680
inside ML-code. They are ``linked'' statically at compile-time, which limits
+ − 681
sometimes their usefulness in cases where, for example, terms needs to be
+ − 682
built up dynamically.
+ − 683
+ − 684
The other kind of antiquotations are \emph{document} antiquotations.
+ − 685
They are used only in the text parts of Isabelle and their purpose is to print
+ − 686
logical entities inside \LaTeX-documents. They are part of the user level and
+ − 687
therefore we are not interested in them in this Tutorial, except in
+ − 688
Appendix \ref{rec:docantiquotations} where we show how to implement your
+ − 689
own document antiquotations.
263
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 690
*}
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 691
15
+ − 692
section {* Terms and Types *}
2
+ − 693
+ − 694
text {*
197
+ − 695
One way to construct Isabelle terms, is by using the antiquotation
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 696
\mbox{@{text "@{term \<dots>}"}}. For example
2
+ − 697
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 698
@{ML_response [display,gray]
75
+ − 699
"@{term \"(a::nat) + b = c\"}"
+ − 700
"Const (\"op =\", \<dots>) $
149
+ − 701
(Const (\"HOL.plus_class.plus\", \<dots>) $ \<dots> $ \<dots>) $ \<dots>"}
2
+ − 702
207
+ − 703
will show the term @{term "(a::nat) + b = c"}, but printed using the internal
256
+ − 704
representation corresponding to the datatype @{ML_type [index] "term"} defined as follows:
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 705
*}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 706
254
+ − 707
ML_val %linenosgray{*datatype term =
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 708
Const of string * typ
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 709
| Free of string * typ
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 710
| Var of indexname * typ
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 711
| Bound of int
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 712
| Abs of string * typ * term
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 713
| $ of term * term *}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 714
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 715
text {*
298
+ − 716
This datatype implements lambda-terms typed in Church-style.
254
+ − 717
As can be seen in Line 5, terms use the usual de Bruijn index mechanism
+ − 718
for representing bound variables. For
251
+ − 719
example in
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 720
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 721
@{ML_response_fake [display, gray]
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 722
"@{term \"\<lambda>x y. x y\"}"
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 723
"Abs (\"x\", \"'a \<Rightarrow> 'b\", Abs (\"y\", \"'a\", Bound 1 $ Bound 0))"}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 724
298
+ − 725
the indices refer to the number of Abstractions (@{ML Abs}) that we need to
+ − 726
skip until we hit the @{ML Abs} that binds the corresponding variable. Constructing
+ − 727
a term with dangling de Bruijn indices is possible, but will be flagged as
+ − 728
ill-formed when you try to typecheck or certify it (see
+ − 729
Section~\ref{sec:typechecking}). Note that the names of bound variables are kept at
+ − 730
abstractions for printing purposes, and so should be treated only as
+ − 731
``comments''. Application in Isabelle is realised with the term-constructor
+ − 732
@{ML $}.
+ − 733
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 734
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 735
Isabelle makes a distinction between \emph{free} variables (term-constructor @{ML Free})
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 736
and variables (term-constructor @{ML Var}). The latter correspond to the schematic
251
+ − 737
variables that when printed show up with a question mark in front of them. Consider
+ − 738
the following two examples
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 739
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 740
@{ML_response_fake [display, gray]
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 741
"let
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 742
val v1 = Var ((\"x\", 3), @{typ bool})
251
+ − 743
val v2 = Var ((\"x1\", 3), @{typ bool})
298
+ − 744
val v3 = Free (\"x\", @{typ bool})
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 745
in
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 746
writeln (Syntax.string_of_term @{context} v1);
298
+ − 747
writeln (Syntax.string_of_term @{context} v2);
+ − 748
writeln (Syntax.string_of_term @{context} v3)
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 749
end"
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 750
"?x3
298
+ − 751
?x1.3
+ − 752
x"}
+ − 753
+ − 754
When constructing terms, you are usually concerned with free variables (as mentioned earlier,
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 755
you cannot construct schematic variables using the antiquotation @{text "@{term \<dots>}"}).
254
+ − 756
If you deal with theorems, you have to, however, observe the distinction. A similar
251
+ − 757
distinction holds for types (see below).
10
+ − 758
2
+ − 759
\begin{readmore}
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 760
Terms and types are described in detail in \isccite{sec:terms}. Their
78
+ − 761
definition and many useful operations are implemented in @{ML_file "Pure/term.ML"}.
254
+ − 762
For constructing terms involving HOL constants, many helper functions are defined
+ − 763
in @{ML_file "HOL/Tools/hologic.ML"}.
2
+ − 764
\end{readmore}
193
+ − 765
+ − 766
Constructing terms via antiquotations has the advantage that only typable
+ − 767
terms can be constructed. For example
2
+ − 768
193
+ − 769
@{ML_response_fake_both [display,gray]
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 770
"@{term \"x x\"}"
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 771
"Type unification failed: Occurs check!"}
193
+ − 772
194
+ − 773
raises a typing error, while it perfectly ok to construct the term
251
+ − 774
+ − 775
@{ML_response_fake [display,gray]
+ − 776
"let
+ − 777
val omega = Free (\"x\", @{typ nat}) $ Free (\"x\", @{typ nat})
+ − 778
in
+ − 779
writeln (Syntax.string_of_term @{context} omega)
+ − 780
end"
+ − 781
"x x"}
193
+ − 782
+ − 783
with the raw ML-constructors.
255
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 784
13
+ − 785
Sometimes the internal representation of terms can be surprisingly different
157
+ − 786
from what you see at the user-level, because the layers of
47
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 787
parsing/type-checking/pretty printing can be quite elaborate.
2
+ − 788
10
+ − 789
\begin{exercise}
2
+ − 790
Look at the internal term representation of the following terms, and
89
+ − 791
find out why they are represented like this:
2
+ − 792
+ − 793
\begin{itemize}
+ − 794
\item @{term "case x of 0 \<Rightarrow> 0 | Suc y \<Rightarrow> y"}
+ − 795
\item @{term "\<lambda>(x,y). P y x"}
+ − 796
\item @{term "{ [x::int] | x. x \<le> -2 }"}
+ − 797
\end{itemize}
+ − 798
+ − 799
Hint: The third term is already quite big, and the pretty printer
+ − 800
may omit parts of it by default. If you want to see all of it, you
122
+ − 801
can use the following ML-function to set the printing depth to a higher
+ − 802
value:
12
+ − 803
75
+ − 804
@{ML [display,gray] "print_depth 50"}
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 805
\end{exercise}
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 806
13
+ − 807
The antiquotation @{text "@{prop \<dots>}"} constructs terms of propositional type,
50
+ − 808
inserting the invisible @{text "Trueprop"}-coercions whenever necessary.
68
+ − 809
Consider for example the pairs
12
+ − 810
126
+ − 811
@{ML_response [display,gray] "(@{term \"P x\"}, @{prop \"P x\"})"
+ − 812
"(Free (\"P\", \<dots>) $ Free (\"x\", \<dots>),
149
+ − 813
Const (\"Trueprop\", \<dots>) $ (Free (\"P\", \<dots>) $ Free (\"x\", \<dots>)))"}
65
+ − 814
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 815
where a coercion is inserted in the second component and
12
+ − 816
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 817
@{ML_response [display,gray] "(@{term \"P x \<Longrightarrow> Q x\"}, @{prop \"P x \<Longrightarrow> Q x\"})"
65
+ − 818
"(Const (\"==>\", \<dots>) $ \<dots> $ \<dots>, Const (\"==>\", \<dots>) $ \<dots> $ \<dots>)"}
12
+ − 819
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 820
where it is not (since it is already constructed by a meta-implication).
19
+ − 821
193
+ − 822
As already seen above, types can be constructed using the antiquotation
+ − 823
@{text "@{typ \<dots>}"}. For example:
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 824
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 825
@{ML_response_fake [display,gray] "@{typ \"bool \<Rightarrow> nat\"}" "bool \<Rightarrow> nat"}
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 826
298
+ − 827
The corresponding datatype is
251
+ − 828
*}
+ − 829
+ − 830
ML_val{*datatype typ =
+ − 831
Type of string * typ list
+ − 832
| TFree of string * sort
+ − 833
| TVar of indexname * sort *}
+ − 834
+ − 835
text {*
+ − 836
Like with terms, there is the distinction between free type
+ − 837
variables (term-constructor @{ML "TFree"} and schematic
+ − 838
type variables (term-constructor @{ML "TVar"}). A type constant,
+ − 839
like @{typ "int"} or @{typ bool}, are types with an empty list
298
+ − 840
of argument types. However, it is a bit difficult to show an
+ − 841
example, because Isabelle always pretty-prints types (unlike terms).
+ − 842
Here is a contrived example:
+ − 843
+ − 844
@{ML_response [display, gray]
+ − 845
"if Type (\"bool\", []) = @{typ \"bool\"} then true else false"
+ − 846
"true"}
251
+ − 847
19
+ − 848
\begin{readmore}
+ − 849
Types are described in detail in \isccite{sec:types}. Their
78
+ − 850
definition and many useful operations are implemented
+ − 851
in @{ML_file "Pure/type.ML"}.
19
+ − 852
\end{readmore}
47
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 853
*}
19
+ − 854
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 855
section {* Constructing Terms and Types Manually\label{sec:terms_types_manually} *}
12
+ − 856
+ − 857
text {*
81
+ − 858
While antiquotations are very convenient for constructing terms, they can
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
+ − 859
only construct fixed terms (remember they are ``linked'' at compile-time).
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 860
However, you often need to construct terms dynamically. For example, a
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 861
function that returns the implication @{text "\<And>(x::nat). P x \<Longrightarrow> Q x"} taking
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 862
@{term P} and @{term Q} as arguments can only be written as:
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
+ − 863
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 864
*}
12
+ − 865
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 866
ML{*fun make_imp P Q =
131
+ − 867
let
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 868
val x = Free ("x", @{typ nat})
131
+ − 869
in
+ − 870
Logic.all x (Logic.mk_implies (P $ x, Q $ x))
+ − 871
end *}
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 872
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 873
text {*
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 874
The reason is that you cannot pass the arguments @{term P} and @{term Q}
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 875
into an antiquotation.\footnote{At least not at the moment.} For example
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 876
the following does \emph{not} work.
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 877
*}
13
+ − 878
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 879
ML{*fun make_wrong_imp P Q = @{prop "\<And>(x::nat). P x \<Longrightarrow> Q x"} *}
12
+ − 880
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 881
text {*
194
+ − 882
To see this, apply @{text "@{term S}"} and @{text "@{term T}"}
193
+ − 883
to both functions. With @{ML make_imp} you obtain the intended term involving
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
+ − 884
the given arguments
65
+ − 885
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 886
@{ML_response [display,gray] "make_imp @{term S} @{term T}"
162
+ − 887
"Const \<dots> $
+ − 888
Abs (\"x\", Type (\"nat\",[]),
+ − 889
Const \<dots> $ (Free (\"S\",\<dots>) $ \<dots>) $ (Free (\"T\",\<dots>) $ \<dots>))"}
68
+ − 890
193
+ − 891
whereas with @{ML make_wrong_imp} you obtain a term involving the @{term "P"}
68
+ − 892
and @{text "Q"} from the antiquotation.
+ − 893
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 894
@{ML_response [display,gray] "make_wrong_imp @{term S} @{term T}"
162
+ − 895
"Const \<dots> $
+ − 896
Abs (\"x\", \<dots>,
+ − 897
Const \<dots> $ (Const \<dots> $ (Free (\"P\",\<dots>) $ \<dots>)) $
192
+ − 898
(Const \<dots> $ (Free (\"Q\",\<dots>) $ \<dots>)))"}
65
+ − 899
192
+ − 900
There are a number of handy functions that are frequently used for
256
+ − 901
constructing terms. One is the function @{ML [index] list_comb}, which takes a term
199
+ − 902
and a list of terms as arguments, and produces as output the term
192
+ − 903
list applied to the term. For example
+ − 904
+ − 905
@{ML_response_fake [display,gray]
298
+ − 906
"let
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 907
val trm = @{term \"P::nat\"}
298
+ − 908
val args = [@{term \"True\"}, @{term \"False\"}]
+ − 909
in
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 910
list_comb (trm, args)
298
+ − 911
end"
192
+ − 912
"Free (\"P\", \"nat\") $ Const (\"True\", \"bool\") $ Const (\"False\", \"bool\")"}
+ − 913
256
+ − 914
Another handy function is @{ML [index] lambda}, which abstracts a variable
193
+ − 915
in a term. For example
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 916
193
+ − 917
@{ML_response_fake [display,gray]
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 918
"let
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 919
val x_nat = @{term \"x::nat\"}
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 920
val trm = @{term \"(P::nat \<Rightarrow> bool) x\"}
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 921
in
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 922
lambda x_nat trm
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 923
end"
193
+ − 924
"Abs (\"x\", \"nat\", Free (\"P\", \"bool \<Rightarrow> bool\") $ Bound 0)"}
+ − 925
298
+ − 926
In this example, @{ML lambda} produces a de Bruijn index (i.e.~@{ML "Bound 0"}),
193
+ − 927
and an abstraction. It also records the type of the abstracted
+ − 928
variable and for printing purposes also its name. Note that because of the
+ − 929
typing annotation on @{text "P"}, the variable @{text "x"} in @{text "P x"}
+ − 930
is of the same type as the abstracted variable. If it is of different type,
+ − 931
as in
+ − 932
+ − 933
@{ML_response_fake [display,gray]
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 934
"let
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 935
val x_int = @{term \"x::int\"}
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 936
val trm = @{term \"(P::nat \<Rightarrow> bool) x\"}
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 937
in
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 938
lambda x_int trm
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 939
end"
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 940
"Abs (\"x\", \"int\", Free (\"P\", \"nat \<Rightarrow> bool\") $ Free (\"x\", \"nat\"))"}
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 941
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 942
then the variable @{text "Free (\"x\", \"int\")"} is \emph{not} abstracted.
193
+ − 943
This is a fundamental principle
+ − 944
of Church-style typing, where variables with the same name still differ, if they
+ − 945
have different type.
192
+ − 946
256
+ − 947
There is also the function @{ML [index] subst_free} with which terms can
194
+ − 948
be replaced by other terms. For example below, we will replace in
292
+ − 949
@{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0 x"}
+ − 950
the subterm @{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0"} by @{term y}, and @{term x} by @{term True}.
49
+ − 951
193
+ − 952
@{ML_response_fake [display,gray]
292
+ − 953
"let
+ − 954
val s1 = (@{term \"(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0\"}, @{term \"y::nat \<Rightarrow> nat\"})
+ − 955
val s2 = (@{term \"x::nat\"}, @{term \"True\"})
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 956
val trm = @{term \"((f::nat \<Rightarrow> nat \<Rightarrow> nat) 0) x\"}
292
+ − 957
in
+ − 958
subst_free [s1, s2] trm
+ − 959
end"
193
+ − 960
"Free (\"y\", \"nat \<Rightarrow> nat\") $ Const (\"True\", \"bool\")"}
+ − 961
+ − 962
As can be seen, @{ML subst_free} does not take typability into account.
+ − 963
However it takes alpha-equivalence into account:
+ − 964
+ − 965
@{ML_response_fake [display, gray]
292
+ − 966
"let
+ − 967
val s = (@{term \"(\<lambda>y::nat. y)\"}, @{term \"x::nat\"})
+ − 968
val trm = @{term \"(\<lambda>x::nat. x)\"}
+ − 969
in
+ − 970
subst_free [s] trm
+ − 971
end"
193
+ − 972
"Free (\"x\", \"nat\")"}
192
+ − 973
255
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 974
There are many convenient functions that construct specific HOL-terms. For
256
+ − 975
example @{ML [index] mk_eq in HOLogic} constructs an equality out of two terms.
255
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 976
The types needed in this equality are calculated from the type of the
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 977
arguments. For example
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 978
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 979
@{ML_response_fake [gray,display]
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 980
"let
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 981
val eq = HOLogic.mk_eq (@{term \"True\"}, @{term \"False\"})
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 982
in
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 983
writeln (Syntax.string_of_term @{context} eq)
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 984
end"
255
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 985
"True = False"}
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 986
*}
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 987
292
+ − 988
text {*
+ − 989
\begin{readmore}
+ − 990
Most of the HOL-specific operations on terms and types are defined
+ − 991
in @{ML_file "HOL/Tools/hologic.ML"}.
+ − 992
\end{readmore}
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 993
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 994
When constructing terms manually, there are a few subtle issues with
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 995
constants. They usually crop up when pattern matching terms or types, or
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 996
when constructing them. While it is perfectly ok to write the function
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 997
@{text is_true} as follows
122
+ − 998
*}
+ − 999
+ − 1000
ML{*fun is_true @{term True} = true
+ − 1001
| is_true _ = false*}
+ − 1002
+ − 1003
text {*
+ − 1004
this does not work for picking out @{text "\<forall>"}-quantified terms. Because
+ − 1005
the function
+ − 1006
*}
+ − 1007
+ − 1008
ML{*fun is_all (@{term All} $ _) = true
+ − 1009
| is_all _ = false*}
+ − 1010
+ − 1011
text {*
275
+ − 1012
will not correctly match the formula @{prop[source] "\<forall>x::nat. P x"}:
122
+ − 1013
+ − 1014
@{ML_response [display,gray] "is_all @{term \"\<forall>x::nat. P x\"}" "false"}
+ − 1015
+ − 1016
The problem is that the @{text "@term"}-antiquotation in the pattern
123
+ − 1017
fixes the type of the constant @{term "All"} to be @{typ "('a \<Rightarrow> bool) \<Rightarrow> bool"} for
122
+ − 1018
an arbitrary, but fixed type @{typ "'a"}. A properly working alternative
+ − 1019
for this function is
+ − 1020
*}
+ − 1021
+ − 1022
ML{*fun is_all (Const ("All", _) $ _) = true
+ − 1023
| is_all _ = false*}
+ − 1024
+ − 1025
text {*
+ − 1026
because now
+ − 1027
+ − 1028
@{ML_response [display,gray] "is_all @{term \"\<forall>x::nat. P x\"}" "true"}
+ − 1029
149
+ − 1030
matches correctly (the first wildcard in the pattern matches any type and the
+ − 1031
second any term).
122
+ − 1032
123
+ − 1033
However there is still a problem: consider the similar function that
131
+ − 1034
attempts to pick out @{text "Nil"}-terms:
122
+ − 1035
*}
+ − 1036
+ − 1037
ML{*fun is_nil (Const ("Nil", _)) = true
+ − 1038
| is_nil _ = false *}
+ − 1039
+ − 1040
text {*
123
+ − 1041
Unfortunately, also this function does \emph{not} work as expected, since
122
+ − 1042
+ − 1043
@{ML_response [display,gray] "is_nil @{term \"Nil\"}" "false"}
+ − 1044
123
+ − 1045
The problem is that on the ML-level the name of a constant is more
149
+ − 1046
subtle than you might expect. The function @{ML is_all} worked correctly,
123
+ − 1047
because @{term "All"} is such a fundamental constant, which can be referenced
+ − 1048
by @{ML "Const (\"All\", some_type)" for some_type}. However, if you look at
122
+ − 1049
+ − 1050
@{ML_response [display,gray] "@{term \"Nil\"}" "Const (\"List.list.Nil\", \<dots>)"}
+ − 1051
131
+ − 1052
the name of the constant @{text "Nil"} depends on the theory in which the
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1053
term constructor is defined (@{text "List"}) and also in which datatype
128
+ − 1054
(@{text "list"}). Even worse, some constants have a name involving
+ − 1055
type-classes. Consider for example the constants for @{term "zero"} and
131
+ − 1056
\mbox{@{text "(op *)"}}:
122
+ − 1057
290
+ − 1058
@{ML_response [display,gray] "(@{term \"0::nat\"}, @{term \"(op *)\"})"
126
+ − 1059
"(Const (\"HOL.zero_class.zero\", \<dots>),
122
+ − 1060
Const (\"HOL.times_class.times\", \<dots>))"}
+ − 1061
123
+ − 1062
While you could use the complete name, for example
+ − 1063
@{ML "Const (\"List.list.Nil\", some_type)" for some_type}, for referring to or
122
+ − 1064
matching against @{text "Nil"}, this would make the code rather brittle.
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1065
The reason is that the theory and the name of the datatype can easily change.
123
+ − 1066
To make the code more robust, it is better to use the antiquotation
122
+ − 1067
@{text "@{const_name \<dots>}"}. With this antiquotation you can harness the
200
+ − 1068
variable parts of the constant's name. Therefore a function for
123
+ − 1069
matching against constants that have a polymorphic type should
+ − 1070
be written as follows.
122
+ − 1071
*}
+ − 1072
+ − 1073
ML{*fun is_nil_or_all (Const (@{const_name "Nil"}, _)) = true
+ − 1074
| is_nil_or_all (Const (@{const_name "All"}, _) $ _) = true
+ − 1075
| is_nil_or_all _ = false *}
+ − 1076
+ − 1077
text {*
274
+ − 1078
\begin{readmore}
+ − 1079
There are many functions in @{ML_file "Pure/term.ML"}, @{ML_file "Pure/logic.ML"} and
+ − 1080
@{ML_file "HOL/Tools/hologic.ML"} that make such manual constructions of terms
+ − 1081
and types easier.\end{readmore}
+ − 1082
+ − 1083
Have a look at these files and try to solve the following two exercises:
+ − 1084
+ − 1085
\begin{exercise}\label{fun:revsum}
+ − 1086
Write a function @{text "rev_sum : term -> term"} that takes a
+ − 1087
term of the form @{text "t\<^isub>1 + t\<^isub>2 + \<dots> + t\<^isub>n"} (whereby @{text "n"} might be one)
+ − 1088
and returns the reversed sum @{text "t\<^isub>n + \<dots> + t\<^isub>2 + t\<^isub>1"}. Assume
+ − 1089
the @{text "t\<^isub>i"} can be arbitrary expressions and also note that @{text "+"}
+ − 1090
associates to the left. Try your function on some examples.
+ − 1091
\end{exercise}
+ − 1092
+ − 1093
\begin{exercise}\label{fun:makesum}
+ − 1094
Write a function which takes two terms representing natural numbers
+ − 1095
in unary notation (like @{term "Suc (Suc (Suc 0))"}), and produces the
+ − 1096
number representing their sum.
+ − 1097
\end{exercise}
+ − 1098
+ − 1099
The antiquotation for properly referencing type constants is @{text "@{type_name \<dots>}"}.
242
+ − 1100
For example
+ − 1101
+ − 1102
@{ML_response [display,gray]
+ − 1103
"@{type_name \"list\"}" "\"List.list\""}
+ − 1104
251
+ − 1105
(FIXME: Explain the following better.)
+ − 1106
200
+ − 1107
Occasionally you have to calculate what the ``base'' name of a given
258
+ − 1108
constant is. For this you can use the function @{ML [index] "Sign.extern_const"} or
+ − 1109
@{ML [index] Long_Name.base_name}. For example:
124
+ − 1110
+ − 1111
@{ML_response [display,gray] "Sign.extern_const @{theory} \"List.list.Nil\"" "\"Nil\""}
122
+ − 1112
124
+ − 1113
The difference between both functions is that @{ML extern_const in Sign} returns
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
+ − 1114
the smallest name that is still unique, whereas @{ML base_name in Long_Name} always
124
+ − 1115
strips off all qualifiers.
122
+ − 1116
+ − 1117
\begin{readmore}
162
+ − 1118
Functions about naming are implemented in @{ML_file "Pure/General/name_space.ML"};
+ − 1119
functions about signatures in @{ML_file "Pure/sign.ML"}.
122
+ − 1120
\end{readmore}
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1121
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1122
Although types of terms can often be inferred, there are many
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1123
situations where you need to construct types manually, especially
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1124
when defining constants. For example the function returning a function
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1125
type is as follows:
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1126
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1127
*}
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1128
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1129
ML{*fun make_fun_type ty1 ty2 = Type ("fun", [ty1, ty2]) *}
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1130
256
+ − 1131
text {* This can be equally written with the combinator @{ML [index] "-->"} as: *}
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1132
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1133
ML{*fun make_fun_type ty1 ty2 = ty1 --> ty2 *}
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1134
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1135
text {*
255
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1136
If you want to construct a function type with more than one argument
256
+ − 1137
type, then you can use @{ML [index] "--->"}.
255
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1138
*}
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1139
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1140
ML{*fun make_fun_types tys ty = tys ---> ty *}
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1141
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1142
text {*
256
+ − 1143
A handy function for manipulating terms is @{ML [index] map_types}: it takes a
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1144
function and applies it to every type in a term. You can, for example,
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1145
change every @{typ nat} in a term into an @{typ int} using the function:
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1146
*}
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1147
252
+ − 1148
ML{*fun nat_to_int ty =
+ − 1149
(case ty of
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1150
@{typ nat} => @{typ int}
252
+ − 1151
| Type (s, tys) => Type (s, map nat_to_int tys)
+ − 1152
| _ => ty)*}
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1153
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1154
text {*
200
+ − 1155
Here is an example:
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1156
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1157
@{ML_response_fake [display,gray]
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1158
"map_types nat_to_int @{term \"a = (1::nat)\"}"
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1159
"Const (\"op =\", \"int \<Rightarrow> int \<Rightarrow> bool\")
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1160
$ Free (\"a\", \"int\") $ Const (\"HOL.one_class.one\", \"int\")"}
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1161
251
+ − 1162
If you want to obtain the list of free type-variables of a term, you
256
+ − 1163
can use the function @{ML [index] add_tfrees in Term}
+ − 1164
(similarly @{ML [index] add_tvars in Term} for the schematic type-variables).
+ − 1165
One would expect that such functions
251
+ − 1166
take a term as input and return a list of types. But their type is actually
+ − 1167
+ − 1168
@{text[display] "Term.term -> (string * Term.sort) list -> (string * Term.sort) list"}
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1169
251
+ − 1170
that is they take, besides a term, also a list of type-variables as input.
+ − 1171
So in order to obtain the list of type-variables of a term you have to
+ − 1172
call them as follows
+ − 1173
+ − 1174
@{ML_response [gray,display]
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1175
"Term.add_tfrees @{term \"(a, b)\"} []"
251
+ − 1176
"[(\"'b\", [\"HOL.type\"]), (\"'a\", [\"HOL.type\"])]"}
+ − 1177
256
+ − 1178
The reason for this definition is that @{ML add_tfrees in Term} can
251
+ − 1179
be easily folded over a list of terms. Similarly for all functions
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1180
named @{text "add_*"} in @{ML_file "Pure/term.ML"}.
251
+ − 1181
122
+ − 1182
*}
162
+ − 1183
122
+ − 1184
298
+ − 1185
section {* Type-Checking\label{sec:typechecking} *}
124
+ − 1186
+ − 1187
text {*
+ − 1188
131
+ − 1189
You can freely construct and manipulate @{ML_type "term"}s and @{ML_type
+ − 1190
typ}es, since they are just arbitrary unchecked trees. However, you
+ − 1191
eventually want to see if a term is well-formed, or type-checks, relative to
256
+ − 1192
a theory. Type-checking is done via the function @{ML [index] cterm_of}, which
+ − 1193
converts a @{ML_type [index] term} into a @{ML_type [index] cterm}, a \emph{certified}
131
+ − 1194
term. Unlike @{ML_type term}s, which are just trees, @{ML_type "cterm"}s are
+ − 1195
abstract objects that are guaranteed to be type-correct, and they can only
+ − 1196
be constructed via ``official interfaces''.
+ − 1197
124
+ − 1198
+ − 1199
Type-checking is always relative to a theory context. For now we use
+ − 1200
the @{ML "@{theory}"} antiquotation to get hold of the current theory.
+ − 1201
For example you can write:
+ − 1202
149
+ − 1203
@{ML_response_fake [display,gray] "cterm_of @{theory} @{term \"(a::nat) + b = c\"}" "a + b = c"}
124
+ − 1204
+ − 1205
This can also be written with an antiquotation:
+ − 1206
+ − 1207
@{ML_response_fake [display,gray] "@{cterm \"(a::nat) + b = c\"}" "a + b = c"}
+ − 1208
+ − 1209
Attempting to obtain the certified term for
+ − 1210
+ − 1211
@{ML_response_fake_both [display,gray] "@{cterm \"1 + True\"}" "Type unification failed \<dots>"}
+ − 1212
+ − 1213
yields an error (since the term is not typable). A slightly more elaborate
+ − 1214
example that type-checks is:
+ − 1215
+ − 1216
@{ML_response_fake [display,gray]
+ − 1217
"let
+ − 1218
val natT = @{typ \"nat\"}
+ − 1219
val zero = @{term \"0::nat\"}
+ − 1220
in
+ − 1221
cterm_of @{theory}
+ − 1222
(Const (@{const_name plus}, natT --> natT --> natT) $ zero $ zero)
+ − 1223
end" "0 + 0"}
+ − 1224
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1225
In Isabelle not just terms need to be certified, but also types. For example,
198
+ − 1226
you obtain the certified type for the Isabelle type @{typ "nat \<Rightarrow> bool"} on
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1227
the ML-level as follows:
160
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1228
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1229
@{ML_response_fake [display,gray]
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1230
"ctyp_of @{theory} (@{typ nat} --> @{typ bool})"
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1231
"nat \<Rightarrow> bool"}
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1232
207
+ − 1233
or with the antiquotation:
+ − 1234
+ − 1235
@{ML_response_fake [display,gray]
+ − 1236
"@{ctyp \"nat \<Rightarrow> bool\"}"
+ − 1237
"nat \<Rightarrow> bool"}
+ − 1238
160
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1239
\begin{readmore}
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1240
For functions related to @{ML_type cterm}s and @{ML_type ctyp}s see
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1241
the file @{ML_file "Pure/thm.ML"}.
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1242
\end{readmore}
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1243
293
+ − 1244
\begin{exercise}
124
+ − 1245
Check that the function defined in Exercise~\ref{fun:revsum} returns a
293
+ − 1246
result that type-checks. See what happens to the solutions of this
+ − 1247
exercise given in \ref{ch:solutions} when they receive an ill-typed term
+ − 1248
as input.
124
+ − 1249
\end{exercise}
+ − 1250
200
+ − 1251
Remember Isabelle follows the Church-style typing for terms, i.e., a term contains
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1252
enough typing information (constants, free variables and abstractions all have typing
124
+ − 1253
information) so that it is always clear what the type of a term is.
256
+ − 1254
Given a well-typed term, the function @{ML [index] type_of} returns the
124
+ − 1255
type of a term. Consider for example:
+ − 1256
+ − 1257
@{ML_response [display,gray]
+ − 1258
"type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
+ − 1259
+ − 1260
To calculate the type, this function traverses the whole term and will
197
+ − 1261
detect any typing inconsistency. For example changing the type of the variable
149
+ − 1262
@{term "x"} from @{typ "nat"} to @{typ "int"} will result in the error message:
124
+ − 1263
+ − 1264
@{ML_response_fake [display,gray]
+ − 1265
"type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})"
+ − 1266
"*** Exception- TYPE (\"type_of: type mismatch in application\" \<dots>"}
+ − 1267
+ − 1268
Since the complete traversal might sometimes be too costly and
256
+ − 1269
not necessary, there is the function @{ML [index] fastype_of}, which
149
+ − 1270
also returns the type of a term.
124
+ − 1271
+ − 1272
@{ML_response [display,gray]
+ − 1273
"fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
+ − 1274
177
+ − 1275
However, efficiency is gained on the expense of skipping some tests. You
124
+ − 1276
can see this in the following example
+ − 1277
+ − 1278
@{ML_response [display,gray]
+ − 1279
"fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})" "bool"}
+ − 1280
149
+ − 1281
where no error is detected.
124
+ − 1282
+ − 1283
Sometimes it is a bit inconvenient to construct a term with
+ − 1284
complete typing annotations, especially in cases where the typing
+ − 1285
information is redundant. A short-cut is to use the ``place-holder''
256
+ − 1286
type @{ML [index] dummyT} and then let type-inference figure out the
124
+ − 1287
complete type. An example is as follows:
+ − 1288
+ − 1289
@{ML_response_fake [display,gray]
+ − 1290
"let
126
+ − 1291
val c = Const (@{const_name \"plus\"}, dummyT)
+ − 1292
val o = @{term \"1::nat\"}
+ − 1293
val v = Free (\"x\", dummyT)
124
+ − 1294
in
126
+ − 1295
Syntax.check_term @{context} (c $ o $ v)
124
+ − 1296
end"
126
+ − 1297
"Const (\"HOL.plus_class.plus\", \"nat \<Rightarrow> nat \<Rightarrow> nat\") $
+ − 1298
Const (\"HOL.one_class.one\", \"nat\") $ Free (\"x\", \"nat\")"}
124
+ − 1299
+ − 1300
Instead of giving explicitly the type for the constant @{text "plus"} and the free
200
+ − 1301
variable @{text "x"}, type-inference fills in the missing information.
124
+ − 1302
+ − 1303
\begin{readmore}
+ − 1304
See @{ML_file "Pure/Syntax/syntax.ML"} where more functions about reading,
200
+ − 1305
checking and pretty-printing of terms are defined. Functions related to
+ − 1306
type-inference are implemented in @{ML_file "Pure/type.ML"} and
160
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1307
@{ML_file "Pure/type_infer.ML"}.
124
+ − 1308
\end{readmore}
162
+ − 1309
+ − 1310
(FIXME: say something about sorts)
124
+ − 1311
*}
+ − 1312
+ − 1313
2
+ − 1314
section {* Theorems *}
+ − 1315
+ − 1316
text {*
50
+ − 1317
Just like @{ML_type cterm}s, theorems are abstract objects of type @{ML_type thm}
201
+ − 1318
that can only be built by going through interfaces. As a consequence, every proof
124
+ − 1319
in Isabelle is correct by construction. This follows the tradition of the LCF approach
+ − 1320
\cite{GordonMilnerWadsworth79}.
107
+ − 1321
2
+ − 1322
78
+ − 1323
To see theorems in ``action'', let us give a proof on the ML-level for the following
+ − 1324
statement:
10
+ − 1325
*}
+ − 1326
+ − 1327
lemma
+ − 1328
assumes assm\<^isub>1: "\<And>(x::nat). P x \<Longrightarrow> Q x"
+ − 1329
and assm\<^isub>2: "P t"
13
+ − 1330
shows "Q t" (*<*)oops(*>*)
10
+ − 1331
+ − 1332
text {*
185
+ − 1333
The corresponding ML-code is as follows:
10
+ − 1334
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1335
@{ML_response_fake [display,gray]
42
+ − 1336
"let
138
+ − 1337
val assm1 = @{cprop \"\<And>(x::nat). P x \<Longrightarrow> Q x\"}
+ − 1338
val assm2 = @{cprop \"(P::nat\<Rightarrow>bool) t\"}
10
+ − 1339
+ − 1340
val Pt_implies_Qt =
+ − 1341
assume assm1
138
+ − 1342
|> forall_elim @{cterm \"t::nat\"};
10
+ − 1343
+ − 1344
val Qt = implies_elim Pt_implies_Qt (assume assm2);
+ − 1345
in
+ − 1346
Qt
+ − 1347
|> implies_intr assm2
+ − 1348
|> implies_intr assm1
48
+ − 1349
end" "\<lbrakk>\<And>x. P x \<Longrightarrow> Q x; P t\<rbrakk> \<Longrightarrow> Q t"}
12
+ − 1350
21
+ − 1351
This code-snippet constructs the following proof:
+ − 1352
+ − 1353
\[
+ − 1354
\infer[(@{text "\<Longrightarrow>"}$-$intro)]{\vdash @{prop "(\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> P t \<Longrightarrow> Q t"}}
+ − 1355
{\infer[(@{text "\<Longrightarrow>"}$-$intro)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"} \vdash @{prop "P t \<Longrightarrow> Q t"}}
+ − 1356
{\infer[(@{text "\<Longrightarrow>"}$-$elim)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"}, @{prop "P t"} \vdash @{prop "Q t"}}
+ − 1357
{\infer[(@{text "\<And>"}$-$elim)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"} \vdash @{prop "P t \<Longrightarrow> Q t"}}
+ − 1358
{\infer[(assume)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"} \vdash @{prop "\<And>x. P x \<Longrightarrow> Q x"}}{}}
+ − 1359
&
+ − 1360
\infer[(assume)]{@{prop "P t"} \vdash @{prop "P t"}}{}
+ − 1361
}
+ − 1362
}
+ − 1363
}
+ − 1364
\]
+ − 1365
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
+ − 1366
However, while we obtained a theorem as result, this theorem is not
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1367
yet stored in Isabelle's theorem database. So it cannot be referenced later
128
+ − 1368
on. How to store theorems will be explained in Section~\ref{sec:storing}.
21
+ − 1369
13
+ − 1370
\begin{readmore}
50
+ − 1371
For the functions @{text "assume"}, @{text "forall_elim"} etc
13
+ − 1372
see \isccite{sec:thms}. The basic functions for theorems are defined in
+ − 1373
@{ML_file "Pure/thm.ML"}.
+ − 1374
\end{readmore}
12
+ − 1375
256
+ − 1376
(FIXME: handy functions working on theorems, like @{ML [index] rulify in ObjectLogic} and so on)
207
+ − 1377
230
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1378
(FIXME: how to add case-names to goal states - maybe in the
207
+ − 1379
next section)
265
+ − 1380
+ − 1381
(FIXME: example for how to add theorem styles)
10
+ − 1382
*}
+ − 1383
265
+ − 1384
ML {*
+ − 1385
fun strip_assums_all (params, Const("all",_) $ Abs(a, T, t)) =
+ − 1386
strip_assums_all ((a, T)::params, t)
+ − 1387
| strip_assums_all (params, B) = (params, B)
+ − 1388
+ − 1389
fun style_parm_premise i ctxt t =
+ − 1390
let val prems = Logic.strip_imp_prems t in
+ − 1391
if i <= length prems
+ − 1392
then let val (params,t) = strip_assums_all([], nth prems (i - 1))
+ − 1393
in subst_bounds(map Free params, t) end
+ − 1394
else error ("Not enough premises for prem" ^ string_of_int i ^
+ − 1395
" in propositon: " ^ Syntax.string_of_term ctxt t)
+ − 1396
end;
+ − 1397
*}
+ − 1398
+ − 1399
ML {*
+ − 1400
strip_assums_all ([], @{term "\<And>x y. A x y"})
+ − 1401
*}
+ − 1402
+ − 1403
setup {*
+ − 1404
TermStyle.add_style "no_all_prem1" (style_parm_premise 1) #>
+ − 1405
TermStyle.add_style "no_all_prem2" (style_parm_premise 2)
+ − 1406
*}
+ − 1407
+ − 1408
243
+ − 1409
section {* Setups (TBD) *}
+ − 1410
+ − 1411
text {*
+ − 1412
In the previous section we used \isacommand{setup} in order to make
+ − 1413
a theorem attribute known to Isabelle. What happens behind the scenes
+ − 1414
is that \isacommand{setup} expects a function of type
+ − 1415
@{ML_type "theory -> theory"}: the input theory is the current theory and the
+ − 1416
output the theory where the theory attribute has been stored.
+ − 1417
+ − 1418
This is a fundamental principle in Isabelle. A similar situation occurs
+ − 1419
for example with declaring constants. The function that declares a
256
+ − 1420
constant on the ML-level is @{ML [index] add_consts_i in Sign}.
243
+ − 1421
If you write\footnote{Recall that ML-code needs to be
+ − 1422
enclosed in \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.}
+ − 1423
*}
+ − 1424
+ − 1425
ML{*Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] @{theory} *}
+ − 1426
+ − 1427
text {*
+ − 1428
for declaring the constant @{text "BAR"} with type @{typ nat} and
+ − 1429
run the code, then you indeed obtain a theory as result. But if you
+ − 1430
query the constant on the Isabelle level using the command \isacommand{term}
+ − 1431
+ − 1432
\begin{isabelle}
+ − 1433
\isacommand{term}~@{text [quotes] "BAR"}\\
+ − 1434
@{text "> \"BAR\" :: \"'a\""}
+ − 1435
\end{isabelle}
+ − 1436
+ − 1437
you do not obtain a constant of type @{typ nat}, but a free variable (printed in
+ − 1438
blue) of polymorphic type. The problem is that the ML-expression above did
+ − 1439
not register the declaration with the current theory. This is what the command
+ − 1440
\isacommand{setup} is for. The constant is properly declared with
+ − 1441
*}
+ − 1442
+ − 1443
setup %gray {* Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] *}
+ − 1444
+ − 1445
text {*
+ − 1446
Now
+ − 1447
+ − 1448
\begin{isabelle}
+ − 1449
\isacommand{term}~@{text [quotes] "BAR"}\\
+ − 1450
@{text "> \"BAR\" :: \"nat\""}
+ − 1451
\end{isabelle}
+ − 1452
+ − 1453
returns a (black) constant with the type @{typ nat}.
+ − 1454
+ − 1455
A similar command is \isacommand{local\_setup}, which expects a function
+ − 1456
of type @{ML_type "local_theory -> local_theory"}. Later on we will also
+ − 1457
use the commands \isacommand{method\_setup} for installing methods in the
+ − 1458
current theory and \isacommand{simproc\_setup} for adding new simprocs to
+ − 1459
the current simpset.
+ − 1460
*}
+ − 1461
292
+ − 1462
section {* Theorem Attributes\label{sec:attributes} *}
123
+ − 1463
127
+ − 1464
text {*
193
+ − 1465
Theorem attributes are @{text "[symmetric]"}, @{text "[THEN \<dots>]"}, @{text
+ − 1466
"[simp]"} and so on. Such attributes are \emph{neither} tags \emph{nor} flags
136
+ − 1467
annotated to theorems, but functions that do further processing once a
193
+ − 1468
theorem is proved. In particular, it is not possible to find out
136
+ − 1469
what are all theorems that have a given attribute in common, unless of course
197
+ − 1470
the function behind the attribute stores the theorems in a retrievable
+ − 1471
data structure.
127
+ − 1472
193
+ − 1473
If you want to print out all currently known attributes a theorem can have,
+ − 1474
you can use the Isabelle command
127
+ − 1475
193
+ − 1476
\begin{isabelle}
+ − 1477
\isacommand{print\_attributes}\\
+ − 1478
@{text "> COMP: direct composition with rules (no lifting)"}\\
+ − 1479
@{text "> HOL.dest: declaration of Classical destruction rule"}\\
+ − 1480
@{text "> HOL.elim: declaration of Classical elimination rule"}\\
+ − 1481
@{text "> \<dots>"}
+ − 1482
\end{isabelle}
+ − 1483
+ − 1484
The theorem attributes fall roughly into two categories: the first category manipulates
207
+ − 1485
the proved theorem (for example @{text "[symmetric]"} and @{text "[THEN \<dots>]"}), and the second
+ − 1486
stores the proved theorem somewhere as data (for example @{text "[simp]"}, which adds
193
+ − 1487
the theorem to the current simpset).
127
+ − 1488
136
+ − 1489
To explain how to write your own attribute, let us start with an extremely simple
+ − 1490
version of the attribute @{text "[symmetric]"}. The purpose of this attribute is
+ − 1491
to produce the ``symmetric'' version of an equation. The main function behind
+ − 1492
this attribute is
127
+ − 1493
*}
+ − 1494
133
+ − 1495
ML{*val my_symmetric = Thm.rule_attribute (fn _ => fn thm => thm RS @{thm sym})*}
+ − 1496
+ − 1497
text {*
256
+ − 1498
where the function @{ML [index] rule_attribute in Thm} expects a function taking a
149
+ − 1499
context (which we ignore in the code above) and a theorem (@{text thm}), and
207
+ − 1500
returns another theorem (namely @{text thm} resolved with the theorem
256
+ − 1501
@{thm [source] sym}: @{thm sym[no_vars]}).\footnote{The function @{ML [index] RS}
+ − 1502
is explained in Section~\ref{sec:simpletacs}.} The function
+ − 1503
@{ML rule_attribute in Thm} then returns an attribute.
136
+ − 1504
+ − 1505
Before we can use the attribute, we need to set it up. This can be done
193
+ − 1506
using the Isabelle command \isacommand{attribute\_setup} as follows:
133
+ − 1507
*}
+ − 1508
193
+ − 1509
attribute_setup %gray my_sym = {* Scan.succeed my_symmetric *}
+ − 1510
"applying the sym rule"
133
+ − 1511
136
+ − 1512
text {*
207
+ − 1513
Inside the @{text "\<verbopen> \<dots> \<verbclose>"}, we have to specify a parser
+ − 1514
for the theorem attribute. Since the attribute does not expect any further
+ − 1515
arguments (unlike @{text "[THEN \<dots>]"}, for example), we use the parser @{ML
+ − 1516
Scan.succeed}. Later on we will also consider attributes taking further
+ − 1517
arguments. An example for the attribute @{text "[my_sym]"} is the proof
136
+ − 1518
*}
+ − 1519
+ − 1520
lemma test[my_sym]: "2 = Suc (Suc 0)" by simp
133
+ − 1521
+ − 1522
text {*
193
+ − 1523
which stores the theorem @{thm test} under the name @{thm [source] test}. You
+ − 1524
can see this, if you query the lemma:
+ − 1525
+ − 1526
\begin{isabelle}
+ − 1527
\isacommand{thm}~@{text "test"}\\
+ − 1528
@{text "> "}~@{thm test}
+ − 1529
\end{isabelle}
+ − 1530
+ − 1531
We can also use the attribute when referring to this theorem:
136
+ − 1532
+ − 1533
\begin{isabelle}
+ − 1534
\isacommand{thm}~@{text "test[my_sym]"}\\
+ − 1535
@{text "> "}~@{thm test[my_sym]}
+ − 1536
\end{isabelle}
+ − 1537
256
+ − 1538
An alternative for setting up an attribute is the function @{ML [index] setup in Attrib}.
243
+ − 1539
So instead of using \isacommand{attribute\_setup}, you can also set up the
+ − 1540
attribute as follows:
+ − 1541
*}
+ − 1542
+ − 1543
ML{*Attrib.setup @{binding "my_sym"} (Scan.succeed my_symmetric)
+ − 1544
"applying the sym rule" *}
+ − 1545
+ − 1546
text {*
+ − 1547
This gives a function from @{ML_type "Context.theory -> Context.theory"}, which
+ − 1548
can be used for example with \isacommand{setup}.
+ − 1549
193
+ − 1550
As an example of a slightly more complicated theorem attribute, we implement
207
+ − 1551
our own version of @{text "[THEN \<dots>]"}. This attribute will take a list of theorems
194
+ − 1552
as argument and resolve the proved theorem with this list (one theorem
+ − 1553
after another). The code for this attribute is
193
+ − 1554
*}
+ − 1555
+ − 1556
ML{*fun MY_THEN thms =
+ − 1557
Thm.rule_attribute (fn _ => fn thm => foldl ((op RS) o swap) thm thms)*}
+ − 1558
+ − 1559
text {*
207
+ − 1560
where @{ML swap} swaps the components of a pair. The setup of this theorem
256
+ − 1561
attribute uses the parser @{ML thms in Attrib}, which parses a list of
193
+ − 1562
theorems.
+ − 1563
*}
+ − 1564
+ − 1565
attribute_setup %gray MY_THEN = {* Attrib.thms >> MY_THEN *}
+ − 1566
"resolving the list of theorems with the proved theorem"
+ − 1567
+ − 1568
text {*
+ − 1569
You can, for example, use this theorem attribute to turn an equation into a
+ − 1570
meta-equation:
+ − 1571
+ − 1572
\begin{isabelle}
+ − 1573
\isacommand{thm}~@{text "test[MY_THEN eq_reflection]"}\\
+ − 1574
@{text "> "}~@{thm test[MY_THEN eq_reflection]}
+ − 1575
\end{isabelle}
+ − 1576
+ − 1577
If you need the symmetric version as a meta-equation, you can write
+ − 1578
+ − 1579
\begin{isabelle}
+ − 1580
\isacommand{thm}~@{text "test[MY_THEN sym eq_reflection]"}\\
+ − 1581
@{text "> "}~@{thm test[MY_THEN sym eq_reflection]}
+ − 1582
\end{isabelle}
+ − 1583
194
+ − 1584
It is also possible to combine different theorem attributes, as in:
193
+ − 1585
+ − 1586
\begin{isabelle}
+ − 1587
\isacommand{thm}~@{text "test[my_sym, MY_THEN eq_reflection]"}\\
+ − 1588
@{text "> "}~@{thm test[my_sym, MY_THEN eq_reflection]}
+ − 1589
\end{isabelle}
+ − 1590
+ − 1591
However, here also a weakness of the concept
194
+ − 1592
of theorem attributes shows through: since theorem attributes can be
193
+ − 1593
arbitrary functions, they do not in general commute. If you try
+ − 1594
+ − 1595
\begin{isabelle}
+ − 1596
\isacommand{thm}~@{text "test[MY_THEN eq_reflection, my_sym]"}\\
+ − 1597
@{text "> "}~@{text "exception THM 1 raised: RSN: no unifiers"}
+ − 1598
\end{isabelle}
+ − 1599
+ − 1600
you get an exception indicating that the theorem @{thm [source] sym}
+ − 1601
does not resolve with meta-equations.
+ − 1602
256
+ − 1603
The purpose of @{ML [index] rule_attribute in Thm} is to directly manipulate theorems.
194
+ − 1604
Another usage of theorem attributes is to add and delete theorems from stored data.
+ − 1605
For example the theorem attribute @{text "[simp]"} adds or deletes a theorem from the
256
+ − 1606
current simpset. For these applications, you can use @{ML [index] declaration_attribute in Thm}.
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1607
To illustrate this function, let us introduce a reference containing a list
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1608
of theorems.
133
+ − 1609
*}
+ − 1610
193
+ − 1611
ML{*val my_thms = ref ([] : thm list)*}
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1612
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1613
text {*
279
+ − 1614
The purpose of this reference is to store a list of theorems.
+ − 1615
We are going to modify it by adding and deleting theorems.
+ − 1616
However, a word of warning: such references must not
207
+ − 1617
be used in any code that is meant to be more than just for testing purposes!
+ − 1618
Here it is only used to illustrate matters. We will show later how to store
+ − 1619
data properly without using references.
193
+ − 1620
207
+ − 1621
We need to provide two functions that add and delete theorems from this list.
153
+ − 1622
For this we use the two functions:
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1623
*}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1624
207
+ − 1625
ML{*fun my_thm_add thm ctxt =
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1626
(my_thms := Thm.add_thm thm (!my_thms); ctxt)
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1627
207
+ − 1628
fun my_thm_del thm ctxt =
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1629
(my_thms := Thm.del_thm thm (!my_thms); ctxt)*}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1630
133
+ − 1631
text {*
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1632
These functions take a theorem and a context and, for what we are explaining
156
+ − 1633
here it is sufficient that they just return the context unchanged. They change
256
+ − 1634
however the reference @{ML my_thms}, whereby the function
+ − 1635
@{ML [index] add_thm in Thm} adds a theorem if it is not already included in
+ − 1636
the list, and @{ML [index] del_thm in Thm} deletes one (both functions use the
+ − 1637
predicate @{ML [index] eq_thm_prop in Thm}, which compares theorems according to
+ − 1638
their proved propositions modulo alpha-equivalence).
133
+ − 1639
207
+ − 1640
You can turn functions @{ML my_thm_add} and @{ML my_thm_del} into
194
+ − 1641
attributes with the code
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1642
*}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1643
207
+ − 1644
ML{*val my_add = Thm.declaration_attribute my_thm_add
+ − 1645
val my_del = Thm.declaration_attribute my_thm_del *}
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1646
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1647
text {*
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1648
and set up the attributes as follows
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1649
*}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1650
193
+ − 1651
attribute_setup %gray my_thms = {* Attrib.add_del my_add my_del *}
207
+ − 1652
"maintaining a list of my_thms - rough test only!"
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1653
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1654
text {*
280
+ − 1655
The parser @{ML [index] add_del in Attrib} is a predefined parser for
194
+ − 1656
adding and deleting lemmas. Now if you prove the next lemma
207
+ − 1657
and attach to it the attribute @{text "[my_thms]"}
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1658
*}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1659
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1660
lemma trueI_2[my_thms]: "True" by simp
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1661
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1662
text {*
194
+ − 1663
then you can see it is added to the initially empty list.
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1664
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1665
@{ML_response_fake [display,gray]
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1666
"!my_thms" "[\"True\"]"}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1667
160
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1668
You can also add theorems using the command \isacommand{declare}.
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1669
*}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1670
207
+ − 1671
declare test[my_thms] trueI_2[my_thms add]
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1672
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1673
text {*
207
+ − 1674
With this attribute, the @{text "add"} operation is the default and does
+ − 1675
not need to be explicitly given. These three declarations will cause the
194
+ − 1676
theorem list to be updated as:
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1677
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1678
@{ML_response_fake [display,gray]
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1679
"!my_thms"
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1680
"[\"True\", \"Suc (Suc 0) = 2\"]"}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1681
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1682
The theorem @{thm [source] trueI_2} only appears once, since the
256
+ − 1683
function @{ML [index] add_thm in Thm} tests for duplicates, before extending
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1684
the list. Deletion from the list works as follows:
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1685
*}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1686
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1687
declare test[my_thms del]
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1688
156
+ − 1689
text {* After this, the theorem list is again:
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1690
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1691
@{ML_response_fake [display,gray]
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1692
"!my_thms"
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1693
"[\"True\"]"}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1694
256
+ − 1695
We used in this example two functions declared as @{ML [index] declaration_attribute in Thm},
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1696
but there can be any number of them. We just have to change the parser for reading
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1697
the arguments accordingly.
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1698
156
+ − 1699
However, as said at the beginning of this example, using references for storing theorems is
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1700
\emph{not} the received way of doing such things. The received way is to
207
+ − 1701
start a ``data slot'', below called @{text MyThmsData}, generated by the functor
194
+ − 1702
@{text GenericDataFun}:
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1703
*}
133
+ − 1704
255
ef1da1abee46
added infrastructure for index; antiquotations have now the options [index] and [indexc]
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1705
ML{*structure MyThmsData = GenericDataFun
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1706
(type T = thm list
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1707
val empty = []
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1708
val extend = I
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1709
fun merge _ = Thm.merge_thms) *}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1710
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1711
text {*
207
+ − 1712
The type @{text "T"} of this data slot is @{ML_type "thm list"}.\footnote{FIXME: give a pointer
+ − 1713
to where data slots are explained properly.}
+ − 1714
To use this data slot, you only have to change @{ML my_thm_add} and
+ − 1715
@{ML my_thm_del} to:
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1716
*}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1717
207
+ − 1718
ML{*val my_thm_add = MyThmsData.map o Thm.add_thm
+ − 1719
val my_thm_del = MyThmsData.map o Thm.del_thm*}
193
+ − 1720
+ − 1721
text {*
194
+ − 1722
where @{ML MyThmsData.map} updates the data appropriately. The
280
+ − 1723
corresponding theorem attributes are
193
+ − 1724
*}
+ − 1725
207
+ − 1726
ML{*val my_add = Thm.declaration_attribute my_thm_add
+ − 1727
val my_del = Thm.declaration_attribute my_thm_del *}
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1728
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1729
text {*
194
+ − 1730
and the setup is as follows
193
+ − 1731
*}
+ − 1732
207
+ − 1733
attribute_setup %gray my_thms2 = {* Attrib.add_del my_add my_del *}
193
+ − 1734
"properly maintaining a list of my_thms"
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1735
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1736
text {*
207
+ − 1737
Initially, the data slot is empty
193
+ − 1738
194
+ − 1739
@{ML_response_fake [display,gray]
+ − 1740
"MyThmsData.get (Context.Proof @{context})"
+ − 1741
"[]"}
193
+ − 1742
194
+ − 1743
but if you prove
+ − 1744
*}
+ − 1745
+ − 1746
lemma three[my_thms2]: "3 = Suc (Suc (Suc 0))" by simp
193
+ − 1747
+ − 1748
text {*
207
+ − 1749
then the lemma is recorded.
194
+ − 1750
+ − 1751
@{ML_response_fake [display,gray]
+ − 1752
"MyThmsData.get (Context.Proof @{context})"
+ − 1753
"[\"3 = Suc (Suc (Suc 0))\"]"}
+ − 1754
207
+ − 1755
With theorem attribute @{text my_thms2} you can also nicely see why it
+ − 1756
is important to
194
+ − 1757
store data in a ``data slot'' and \emph{not} in a reference. Backtrack
207
+ − 1758
to the point just before the lemma @{thm [source] three} was proved and
256
+ − 1759
check the the content of @{ML_struct MyThmsData}: it should be empty.
207
+ − 1760
The addition has been properly retracted. Now consider the proof:
194
+ − 1761
*}
+ − 1762
+ − 1763
lemma four[my_thms]: "4 = Suc (Suc (Suc (Suc 0)))" by simp
193
+ − 1764
194
+ − 1765
text {*
+ − 1766
Checking the content of @{ML my_thms} gives
+ − 1767
+ − 1768
@{ML_response_fake [display,gray]
+ − 1769
"!my_thms"
+ − 1770
"[\"4 = Suc (Suc (Suc (Suc 0)))\", \"True\"]"}
+ − 1771
207
+ − 1772
as expected, but if you backtrack before the lemma @{thm [source] four}, the
194
+ − 1773
content of @{ML my_thms} is unchanged. The backtracking mechanism
207
+ − 1774
of Isabelle is completely oblivious about what to do with references, but
+ − 1775
properly treats ``data slots''!
194
+ − 1776
207
+ − 1777
Since storing theorems in a list is such a common task, there is the special
265
+ − 1778
functor @{ML_functor Named_Thms}, which does most of the work for you. To obtain
281
+ − 1779
a named theorem list, you just declare
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1780
*}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1781
265
+ − 1782
ML{*structure FooRules = Named_Thms
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1783
(val name = "foo"
194
+ − 1784
val description = "Rules for foo") *}
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1785
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1786
text {*
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1787
and set up the @{ML_struct FooRules} with the command
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1788
*}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1789
177
+ − 1790
setup %gray {* FooRules.setup *}
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1791
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1792
text {*
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1793
This code declares a data slot where the theorems are stored,
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1794
an attribute @{text foo} (with the @{text add} and @{text del} options
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1795
for adding and deleting theorems) and an internal ML interface to retrieve and
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1796
modify the theorems.
133
+ − 1797
157
+ − 1798
Furthermore, the facts are made available on the user-level under the dynamic
156
+ − 1799
fact name @{text foo}. For example you can declare three lemmas to be of the kind
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1800
@{text foo} by:
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1801
*}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1802
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1803
lemma rule1[foo]: "A" sorry
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1804
lemma rule2[foo]: "B" sorry
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1805
lemma rule3[foo]: "C" sorry
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1806
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1807
text {* and undeclare the first one by: *}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1808
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1809
declare rule1[foo del]
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1810
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1811
text {* and query the remaining ones with:
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1812
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1813
\begin{isabelle}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1814
\isacommand{thm}~@{text "foo"}\\
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1815
@{text "> ?C"}\\
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1816
@{text "> ?B"}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1817
\end{isabelle}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1818
156
+ − 1819
On the ML-level the rules marked with @{text "foo"} can be retrieved
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1820
using the function @{ML FooRules.get}:
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1821
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1822
@{ML_response_fake [display,gray] "FooRules.get @{context}" "[\"?C\",\"?B\"]"}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1823
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1824
\begin{readmore}
263
195c4444dff7
added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1825
For more information see @{ML_file "Pure/Tools/named_thms.ML"}.
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1826
\end{readmore}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1827
156
+ − 1828
(FIXME What are: @{text "theory_attributes"}, @{text "proof_attributes"}?)
133
+ − 1829
+ − 1830
+ − 1831
\begin{readmore}
207
+ − 1832
FIXME: @{ML_file "Pure/more_thm.ML"}; parsers for attributes is in
+ − 1833
@{ML_file "Pure/Isar/attrib.ML"}...also explained in the chapter about
+ − 1834
parsing.
133
+ − 1835
\end{readmore}
+ − 1836
*}
+ − 1837
127
+ − 1838
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1839
153
+ − 1840
section {* Theories, Contexts and Local Theories (TBD) *}
123
+ − 1841
126
+ − 1842
text {*
+ − 1843
There are theories, proof contexts and local theories (in this order, if you
+ − 1844
want to order them).
+ − 1845
+ − 1846
In contrast to an ordinary theory, which simply consists of a type
+ − 1847
signature, as well as tables for constants, axioms and theorems, a local
202
+ − 1848
theory contains additional context information, such as locally fixed
126
+ − 1849
variables and local assumptions that may be used by the package. The type
+ − 1850
@{ML_type local_theory} is identical to the type of \emph{proof contexts}
+ − 1851
@{ML_type "Proof.context"}, although not every proof context constitutes a
+ − 1852
valid local theory.
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1853
*}
126
+ − 1854
235
+ − 1855
(*
229
+ − 1856
ML{*signature UNIVERSAL_TYPE =
+ − 1857
sig
+ − 1858
type t
+ − 1859
+ − 1860
val embed: unit -> ('a -> t) * (t -> 'a option)
+ − 1861
end*}
+ − 1862
+ − 1863
ML{*structure U:> UNIVERSAL_TYPE =
+ − 1864
struct
+ − 1865
type t = exn
+ − 1866
+ − 1867
fun 'a embed () =
+ − 1868
let
+ − 1869
exception E of 'a
+ − 1870
fun project (e: t): 'a option =
+ − 1871
case e of
+ − 1872
E a => SOME a
+ − 1873
| _ => NONE
+ − 1874
in
+ − 1875
(E, project)
+ − 1876
end
+ − 1877
end*}
+ − 1878
+ − 1879
text {*
+ − 1880
The idea is that type t is the universal type and that each call to embed
+ − 1881
returns a new pair of functions (inject, project), where inject embeds a
+ − 1882
value into the universal type and project extracts the value from the
+ − 1883
universal type. A pair (inject, project) returned by embed works together in
+ − 1884
that project u will return SOME v if and only if u was created by inject
+ − 1885
v. If u was created by a different function inject', then project returns
+ − 1886
NONE.
+ − 1887
+ − 1888
in library.ML
+ − 1889
*}
+ − 1890
+ − 1891
ML_val{*structure Object = struct type T = exn end; *}
+ − 1892
+ − 1893
ML{*functor Test (U: UNIVERSAL_TYPE): sig end =
+ − 1894
struct
+ − 1895
val (intIn: int -> U.t, intOut) = U.embed ()
+ − 1896
val r: U.t ref = ref (intIn 13)
+ − 1897
val s1 =
+ − 1898
case intOut (!r) of
+ − 1899
NONE => "NONE"
+ − 1900
| SOME i => Int.toString i
+ − 1901
val (realIn: real -> U.t, realOut) = U.embed ()
+ − 1902
val () = r := realIn 13.0
+ − 1903
val s2 =
+ − 1904
case intOut (!r) of
+ − 1905
NONE => "NONE"
+ − 1906
| SOME i => Int.toString i
+ − 1907
val s3 =
+ − 1908
case realOut (!r) of
+ − 1909
NONE => "NONE"
+ − 1910
| SOME x => Real.toString x
239
+ − 1911
val () = writeln (concat [s1, " ", s2, " ", s3, "\n"])
229
+ − 1912
end*}
+ − 1913
+ − 1914
ML_val{*structure t = Test(U) *}
+ − 1915
+ − 1916
ML_val{*structure Datatab = TableFun(type key = int val ord = int_ord);*}
+ − 1917
235
+ − 1918
ML {* LocalTheory.restore *}
+ − 1919
ML {* LocalTheory.set_group *}
+ − 1920
*)
+ − 1921
153
+ − 1922
section {* Storing Theorems\label{sec:storing} (TBD) *}
123
+ − 1923
256
+ − 1924
text {* @{ML [index] add_thms_dynamic in PureThy} *}
123
+ − 1925
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1926
local_setup {*
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1927
LocalTheory.note Thm.theoremK
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1928
((@{binding "allI_alt"}, []), [@{thm allI}]) #> snd *}
100
+ − 1929
75
+ − 1930
126
+ − 1931
(* FIXME: some code below *)
89
+ − 1932
+ − 1933
(*<*)
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
+ − 1934
(*
89
+ − 1935
setup {*
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
+ − 1936
Sign.add_consts_i [(Binding"bar", @{typ "nat"},NoSyn)]
89
+ − 1937
*}
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
+ − 1938
*)
89
+ − 1939
lemma "bar = (1::nat)"
+ − 1940
oops
+ − 1941
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
+ − 1942
(*
89
+ − 1943
setup {*
+ − 1944
Sign.add_consts_i [("foo", @{typ "nat"},NoSyn)]
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1945
#> PureThy.add_defs false [((@{binding "foo_def"},
89
+ − 1946
Logic.mk_equals (Const ("FirstSteps.foo", @{typ "nat"}), @{term "1::nat"})), [])]
+ − 1947
#> snd
+ − 1948
*}
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
+ − 1949
*)
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1950
(*
89
+ − 1951
lemma "foo = (1::nat)"
+ − 1952
apply(simp add: foo_def)
+ − 1953
done
+ − 1954
+ − 1955
thm foo_def
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
+ − 1956
*)
89
+ − 1957
(*>*)
+ − 1958
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1959
section {* Pretty-Printing\label{sec:pretty} *}
247
+ − 1960
+ − 1961
text {*
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1962
So far we printed out only plain strings without any formatting except for
280
+ − 1963
occasional explicit line breaks using @{text [quotes] "\\n"}. This is
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1964
sufficient for ``quick-and-dirty'' printouts. For something more
249
+ − 1965
sophisticated, Isabelle includes an infrastructure for properly formatting text.
+ − 1966
This infrastructure is loosely based on a paper by Oppen~\cite{Oppen80}. Most of
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1967
its functions do not operate on @{ML_type string}s, but on instances of the
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1968
type:
247
+ − 1969
256
+ − 1970
@{ML_type [display, gray, index] "Pretty.T"}
+ − 1971
+ − 1972
The function @{ML str in Pretty} transforms a (plain) string into such a pretty
247
+ − 1973
type. For example
+ − 1974
+ − 1975
@{ML_response_fake [display,gray]
+ − 1976
"Pretty.str \"test\"" "String (\"test\", 4)"}
+ − 1977
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1978
where the result indicates that we transformed a string with length 4. Once
249
+ − 1979
you have a pretty type, you can, for example, control where linebreaks may
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1980
occur in case the text wraps over a line, or with how much indentation a
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1981
text should be printed. However, if you want to actually output the
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1982
formatted text, you have to transform the pretty type back into a @{ML_type
256
+ − 1983
string}. This can be done with the function @{ML [index] string_of in Pretty}. In what
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1984
follows we will use the following wrapper function for printing a pretty
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1985
type:
247
+ − 1986
*}
+ − 1987
+ − 1988
ML{*fun pprint prt = writeln (Pretty.string_of prt)*}
153
+ − 1989
+ − 1990
text {*
249
+ − 1991
The point of the pretty-printing infrastructure is to give hints about how to
+ − 1992
layout text and let Isabelle do the actual layout. Let us first explain
280
+ − 1993
how you can insert places where a line break can occur. For this assume the
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1994
following function that replicates a string n times:
247
+ − 1995
*}
+ − 1996
+ − 1997
ML{*fun rep n str = implode (replicate n str) *}
+ − 1998
+ − 1999
text {*
+ − 2000
and suppose we want to print out the string:
+ − 2001
*}
+ − 2002
+ − 2003
ML{*val test_str = rep 8 "fooooooooooooooobaaaaaaaaaaaar "*}
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2004
247
+ − 2005
text {*
282
+ − 2006
We deliberately chose a large string so that it spans over more than one line.
249
+ − 2007
If we print out the string using the usual ``quick-and-dirty'' method, then
+ − 2008
we obtain the ugly output:
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2009
247
+ − 2010
@{ML_response_fake [display,gray]
+ − 2011
"writeln test_str"
+ − 2012
"fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar foooooooooo
+ − 2013
ooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaa
+ − 2014
aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo
+ − 2015
oooooooooooooobaaaaaaaaaaaar"}
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2016
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2017
We obtain the same if we just use
247
+ − 2018
+ − 2019
@{ML_response_fake [display,gray]
+ − 2020
"pprint (Pretty.str test_str)"
+ − 2021
"fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar foooooooooo
+ − 2022
ooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaa
+ − 2023
aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo
+ − 2024
oooooooooooooobaaaaaaaaaaaar"}
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2025
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2026
However by using pretty types you have the ability to indicate a possible
280
+ − 2027
line break for example at each space. You can achieve this with the function
256
+ − 2028
@{ML [index] breaks in Pretty}, which expects a list of pretty types and inserts a
280
+ − 2029
possible line break in between every two elements in this list. To print
249
+ − 2030
this list of pretty types as a single string, we concatenate them
256
+ − 2031
with the function @{ML [index] blk in Pretty} as follows:
210
+ − 2032
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2033
247
+ − 2034
@{ML_response_fake [display,gray]
+ − 2035
"let
+ − 2036
val ptrs = map Pretty.str (space_explode \" \" test_str)
+ − 2037
in
+ − 2038
pprint (Pretty.blk (0, Pretty.breaks ptrs))
+ − 2039
end"
+ − 2040
"fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
+ − 2041
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
+ − 2042
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
+ − 2043
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2044
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2045
Here the layout of @{ML test_str} is much more pleasing to the
256
+ − 2046
eye. The @{ML "0"} in @{ML [index] blk in Pretty} stands for no indentation
249
+ − 2047
of the printed string. You can increase the indentation and obtain
247
+ − 2048
+ − 2049
@{ML_response_fake [display,gray]
+ − 2050
"let
+ − 2051
val ptrs = map Pretty.str (space_explode \" \" test_str)
+ − 2052
in
+ − 2053
pprint (Pretty.blk (3, Pretty.breaks ptrs))
+ − 2054
end"
+ − 2055
"fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
+ − 2056
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
+ − 2057
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
+ − 2058
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2059
247
+ − 2060
where starting from the second line the indent is 3. If you want
+ − 2061
that every line starts with the same indent, you can use the
256
+ − 2062
function @{ML [index] indent in Pretty} as follows:
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2063
247
+ − 2064
@{ML_response_fake [display,gray]
+ − 2065
"let
+ − 2066
val ptrs = map Pretty.str (space_explode \" \" test_str)
+ − 2067
in
+ − 2068
pprint (Pretty.indent 10 (Pretty.blk (0, Pretty.breaks ptrs)))
+ − 2069
end"
+ − 2070
" fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
+ − 2071
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
+ − 2072
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
+ − 2073
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2074
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2075
If you want to print out a list of items separated by commas and
249
+ − 2076
have the linebreaks handled properly, you can use the function
256
+ − 2077
@{ML [index] commas in Pretty}. For example
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2078
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2079
@{ML_response_fake [display,gray]
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2080
"let
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2081
val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020)
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2082
in
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2083
pprint (Pretty.blk (0, Pretty.commas ptrs))
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2084
end"
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2085
"99998, 99999, 100000, 100001, 100002, 100003, 100004, 100005, 100006,
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2086
100007, 100008, 100009, 100010, 100011, 100012, 100013, 100014, 100015,
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2087
100016, 100017, 100018, 100019, 100020"}
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2088
257
+ − 2089
where @{ML upto} generates a list of integers. You can print out this
283
+ − 2090
list as a ``set'', that means enclosed inside @{text [quotes] "{"} and
249
+ − 2091
@{text [quotes] "}"}, and separated by commas using the function
256
+ − 2092
@{ML [index] enum in Pretty}. For example
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2093
*}
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2094
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2095
text {*
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2096
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2097
@{ML_response_fake [display,gray]
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2098
"let
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2099
val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020)
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2100
in
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2101
pprint (Pretty.enum \",\" \"{\" \"}\" ptrs)
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2102
end"
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2103
"{99998, 99999, 100000, 100001, 100002, 100003, 100004, 100005, 100006,
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2104
100007, 100008, 100009, 100010, 100011, 100012, 100013, 100014, 100015,
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2105
100016, 100017, 100018, 100019, 100020}"}
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2106
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2107
As can be seen, this function prints out the ``set'' so that starting
257
+ − 2108
from the second, each new line as an indentation of 2.
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2109
257
+ − 2110
If you print out something that goes beyond the capabilities of the
280
+ − 2111
standard functions, you can do relatively easily the formatting
249
+ − 2112
yourself. Assume you want to print out a list of items where like in ``English''
+ − 2113
the last two items are separated by @{text [quotes] "and"}. For this you can
+ − 2114
write the function
+ − 2115
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2116
*}
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2117
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2118
ML %linenosgray{*fun and_list [] = []
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2119
| and_list [x] = [x]
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2120
| and_list xs =
247
+ − 2121
let
+ − 2122
val (front, last) = split_last xs
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2123
in
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2124
(Pretty.commas front) @
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2125
[Pretty.brk 1, Pretty.str "and", Pretty.brk 1, last]
247
+ − 2126
end *}
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2127
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2128
text {*
257
+ − 2129
where Line 7 prints the beginning of the list and Line
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2130
8 the last item. We have to use @{ML "Pretty.brk 1"} in order
249
+ − 2131
to insert a space (of length 1) before the
280
+ − 2132
@{text [quotes] "and"}. This space is also a place where a line break
+ − 2133
can occur. We do the same after the @{text [quotes] "and"}. This gives you
249
+ − 2134
for example
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2135
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2136
@{ML_response_fake [display,gray]
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2137
"let
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2138
val ptrs = map (Pretty.str o string_of_int) (1 upto 22)
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2139
in
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2140
pprint (Pretty.blk (0, and_list ptrs))
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2141
end"
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2142
"1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2143
and 22"}
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2144
249
+ − 2145
Next we like to pretty-print a term and its type. For this we use the
+ − 2146
function @{text "tell_type"}.
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2147
*}
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2148
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2149
ML %linenosgray{*fun tell_type ctxt t =
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2150
let
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2151
fun pstr s = Pretty.breaks (map Pretty.str (space_explode " " s))
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2152
val ptrm = Pretty.quote (Syntax.pretty_term ctxt t)
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2153
val pty = Pretty.quote (Syntax.pretty_typ ctxt (fastype_of t))
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2154
in
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2155
pprint (Pretty.blk (0,
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2156
(pstr "The term " @ [ptrm] @ pstr " has type "
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2157
@ [pty, Pretty.str "."])))
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2158
end*}
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2159
210
+ − 2160
text {*
257
+ − 2161
In Line 3 we define a function that inserts possible linebreaks in places
+ − 2162
where a space is. In Lines 4 and 5 we pretty-print the term and its type
+ − 2163
using the functions @{ML [index] pretty_term in Syntax} and @{ML [index]
+ − 2164
pretty_typ in Syntax}. We also use the function @{ML [index] quote in
+ − 2165
Pretty} in order to enclose the term and type inside quotation marks. In
+ − 2166
Line 9 we add a period right after the type without the possibility of a
280
+ − 2167
line break, because we do not want that a line break occurs there.
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2168
257
+ − 2169
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2170
Now you can write
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2171
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2172
@{ML_response_fake [display,gray]
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2173
"tell_type @{context} @{term \"min (Suc 0)\"}"
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2174
"The term \"min (Suc 0)\" has type \"nat \<Rightarrow> nat\"."}
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2175
280
+ − 2176
To see the proper line breaking, you can try out the function on a bigger term
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2177
and type. For example:
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2178
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2179
@{ML_response_fake [display,gray]
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2180
"tell_type @{context} @{term \"op = (op = (op = (op = (op = op =))))\"}"
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2181
"The term \"op = (op = (op = (op = (op = op =))))\" has type
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2182
\"((((('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool) \<Rightarrow> bool) \<Rightarrow> bool) \<Rightarrow> bool) \<Rightarrow> bool\"."}
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2183
257
+ − 2184
+ − 2185
FIXME: TBD below
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2186
*}
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2187
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2188
ML {* pprint (Pretty.big_list "header" (map (Pretty.str o string_of_int) (4 upto 10))) *}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2189
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2190
text {*
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2191
chunks inserts forced breaks (unlike blk where you have to do this yourself)
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2192
*}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2193
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2194
ML {* (Pretty.chunks [Pretty.str "a", Pretty.str "b"],
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2195
Pretty.blk (0, [Pretty.str "a", Pretty.str "b"])) *}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2196
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2197
ML {*
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2198
fun setmp_show_all_types f =
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2199
setmp show_all_types
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2200
(! show_types orelse ! show_sorts orelse ! show_all_types) f;
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2201
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2202
val ctxt = @{context};
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2203
val t1 = @{term "undefined::nat"};
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2204
val t2 = @{term "Suc y"};
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2205
val pty = Pretty.block (Pretty.breaks
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2206
[(setmp show_question_marks false o setmp_show_all_types)
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2207
(Syntax.pretty_term ctxt) t1,
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2208
Pretty.str "=", Syntax.pretty_term ctxt t2]);
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2209
pty |> Pretty.string_of |> priority
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2210
*}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2211
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2212
text {* the infrastructure of Syntax-pretty-term makes sure it is printed nicely *}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2213
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2214
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2215
ML {* Pretty.mark Markup.hilite (Pretty.str "foo") |> Pretty.string_of |> writeln *}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2216
ML {* (Pretty.str "bar") |> Pretty.string_of |> writeln *}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2217
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2218
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2219
ML {* Pretty.mark Markup.subgoal (Pretty.str "foo") |> Pretty.string_of |> writeln *}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2220
ML {* (Pretty.str "bar") |> Pretty.string_of |> writeln *}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2221
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2222
text {*
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2223
Still to be done:
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2224
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2225
What happens with big formulae?
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2226
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2227
\begin{readmore}
249
+ − 2228
The general infrastructure for pretty-printing is implemented in the file
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2229
@{ML_file "Pure/General/pretty.ML"}. The file @{ML_file "Pure/Syntax/syntax.ML"}
249
+ − 2230
contains pretty-printing functions for terms, types, theorems and so on.
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2231
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2232
@{ML_file "Pure/General/markup.ML"}
248
11851b20fb78
added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2233
\end{readmore}
153
+ − 2234
*}
+ − 2235
250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2236
text {*
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2237
printing into the goal buffer as part of the proof state
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2238
*}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2239
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2240
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2241
ML {* Pretty.mark Markup.hilite (Pretty.str "foo") |> Pretty.string_of |> writeln *}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2242
ML {* (Pretty.str "bar") |> Pretty.string_of |> writeln *}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2243
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2244
text {* writing into the goal buffer *}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2245
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2246
ML {* fun my_hook interactive state =
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2247
(interactive ? Proof.goal_message (fn () => Pretty.str
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2248
"foo")) state
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2249
*}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2250
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2251
setup {* Context.theory_map (Specification.add_theorem_hook my_hook) *}
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2252
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2253
lemma "False"
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2254
oops
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2255
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2256
153
+ − 2257
section {* Misc (TBD) *}
92
+ − 2258
262
+ − 2259
ML {*Datatype.get_info @{theory} "List.list"*}
92
+ − 2260
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2261
section {* Name Space(TBD) *}
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2262
196
+ − 2263
end