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