2
+ − 1
theory FirstSteps
25
+ − 2
imports Base
2
+ − 3
begin
+ − 4
+ − 5
chapter {* First Steps *}
+ − 6
42
+ − 7
text {*
+ − 8
54
+ − 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
6
+ − 13
\begin{center}
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}
6
+ − 20
\end{center}
157
+ − 21
+ − 22
We also generally assume you are working with HOL. The given examples might
+ − 23
need to be adapted slightly 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
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
+ − 41
Like normal Isabelle proof scripts, \isacommand{ML}-commands can be
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
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 44
value and function bindings, and even those can be undone when the proof
160
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 45
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
+ − 46
\isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"} scaffolding whenever we
107
+ − 47
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
+ − 48
code, rather they indicate what the response is when the code is evaluated.
10
+ − 49
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
+ − 50
Once a portion of code is relatively stable, you usually want to export it
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 51
to a separate ML-file. Such files can then be included in a theory by using
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 52
the \isacommand{uses}-command in the header of the theory, like:
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 53
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 54
\begin{center}
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 55
\begin{tabular}{@ {}l}
54
+ − 56
\isacommand{theory} FirstSteps\\
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 57
\isacommand{imports} Main\\
58
+ − 58
\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
+ − 59
\isacommand{begin}\\
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 60
\ldots
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 61
\end{tabular}
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 62
\end{center}
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 63
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 64
*}
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 65
126
+ − 66
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
+ − 67
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 68
text {*
13
+ − 69
50
+ − 70
During development you might find it necessary to inspect some data
10
+ − 71
in your code. This can be done in a ``quick-and-dirty'' fashion using
12
+ − 72
the function @{ML "warning"}. For example
10
+ − 73
161
+ − 74
@{ML_response_fake [display,gray] "warning \"any string\"" "\"any string\""}
10
+ − 75
58
+ − 76
will print out @{text [quotes] "any string"} inside the response buffer
+ − 77
of Isabelle. This function expects a string as argument. If you develop under PolyML,
50
+ − 78
then there is a convenient, though again ``quick-and-dirty'', method for
107
+ − 79
converting values into strings, namely the function @{ML makestring}:
10
+ − 80
161
+ − 81
@{ML_response_fake [display,gray] "warning (makestring 1)" "\"1\""}
12
+ − 82
65
+ − 83
However @{ML makestring} only works if the type of what is converted is monomorphic
78
+ − 84
and not a function.
12
+ − 85
52
+ − 86
The function @{ML "warning"} should only be used for testing purposes, because any
+ − 87
output this function generates will be overwritten as soon as an error is
50
+ − 88
raised. For printing anything more serious and elaborate, the
54
+ − 89
function @{ML tracing} is more appropriate. This function writes all output into
89
+ − 90
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
+ − 91
161
+ − 92
@{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
+ − 93
58
+ − 94
It is also possible to redirect the ``channel'' where the string @{text "foo"} is
78
+ − 95
printed to a separate file, e.g.~to prevent ProofGeneral from choking on massive
107
+ − 96
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
+ − 97
*}
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 98
69
+ − 99
ML{*val strip_specials =
42
+ − 100
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
+ − 101
fun strip ("\^A" :: _ :: cs) = strip cs
42
+ − 102
| strip (c :: cs) = c :: strip cs
+ − 103
| strip [] = [];
+ − 104
in implode o strip o explode end;
+ − 105
+ − 106
fun redirect_tracing stream =
+ − 107
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
+ − 108
(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
+ − 109
TextIO.output (stream, "\n");
69
+ − 110
TextIO.flushOut stream)) *}
14
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 111
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 112
text {*
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 113
Calling @{ML "redirect_tracing"} with @{ML "(TextIO.openOut \"foo.bar\")"}
58
+ − 114
will cause that all tracing information is printed into the file @{text "foo.bar"}.
75
+ − 115
107
+ − 116
You can print out error messages with the function @{ML error}; for example:
75
+ − 117
122
+ − 118
@{ML_response_fake [display,gray] "if 0=1 then true else (error \"foo\")"
+ − 119
"Exception- ERROR \"foo\" raised
+ − 120
At command \"ML\"."}
75
+ − 121
192
+ − 122
(FIXME @{ML Toplevel.debug} @{ML Toplevel.profiling})
+ − 123
*}
+ − 124
+ − 125
(*
+ − 126
ML {*
+ − 127
fun dodgy_fun () = (raise (ERROR "bar"); 1)
+ − 128
*}
+ − 129
+ − 130
ML {* set Toplevel.debug *}
+ − 131
+ − 132
ML {* fun f1 () = dodgy_fun () *}
+ − 133
+ − 134
ML {* f1 () *}
+ − 135
*)
+ − 136
+ − 137
text {*
126
+ − 138
Most often you want to inspect data of type @{ML_type term}, @{ML_type cterm}
+ − 139
or @{ML_type thm}. Isabelle contains elaborate pretty-printing functions for printing them,
+ − 140
but for quick-and-dirty solutions they are far too unwieldy. A simple way to transform
+ − 141
a term into a string is to use the function @{ML Syntax.string_of_term}.
+ − 142
+ − 143
@{ML_response_fake [display,gray]
+ − 144
"Syntax.string_of_term @{context} @{term \"1::nat\"}"
+ − 145
"\"\\^E\\^Fterm\\^E\\^E\\^Fconst\\^Fname=HOL.one_class.one\\^E1\\^E\\^F\\^E\\^E\\^F\\^E\""}
+ − 146
+ − 147
This produces a string with some additional information encoded in it. The string
+ − 148
can be properly printed by using the function @{ML warning}.
+ − 149
+ − 150
@{ML_response_fake [display,gray]
+ − 151
"warning (Syntax.string_of_term @{context} @{term \"1::nat\"})"
+ − 152
"\"1\""}
+ − 153
+ − 154
A @{ML_type cterm} can be transformed into a string by the following function.
+ − 155
*}
+ − 156
+ − 157
ML{*fun str_of_cterm ctxt t =
+ − 158
Syntax.string_of_term ctxt (term_of t)*}
+ − 159
+ − 160
text {*
149
+ − 161
In this example the function @{ML term_of} extracts the @{ML_type term} from
+ − 162
a @{ML_type cterm}. If there are more than one @{ML_type cterm}s to be
+ − 163
printed, you can use the function @{ML commas} to separate them.
126
+ − 164
*}
+ − 165
+ − 166
ML{*fun str_of_cterms ctxt ts =
+ − 167
commas (map (str_of_cterm ctxt) ts)*}
+ − 168
+ − 169
text {*
+ − 170
The easiest way to get the string of a theorem is to transform it
190
+ − 171
into a @{ML_type cterm} using the function @{ML crep_thm}.
+ − 172
*}
+ − 173
+ − 174
ML{*fun str_of_thm_raw ctxt thm =
+ − 175
str_of_cterm ctxt (#prop (crep_thm thm))*}
+ − 176
+ − 177
text {*
+ − 178
Theorems also include schematic variables, such as @{text "?P"},
+ − 179
@{text "?Q"} and so on.
+ − 180
+ − 181
@{ML_response_fake [display, gray]
+ − 182
"warning (str_of_thm_raw @{context} @{thm conjI})"
+ − 183
"\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q"}
+ − 184
+ − 185
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
+ − 186
these schematic variables into free variables using the
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 187
function @{ML Variable.import_thms}.
126
+ − 188
*}
+ − 189
158
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 190
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
+ − 191
let
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 192
val ((_, [thm']), _) = Variable.import_thms true [thm] ctxt
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 193
in
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 194
thm'
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 195
end
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 196
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 197
fun str_of_thm ctxt thm =
171
+ − 198
str_of_cterm ctxt (#prop (crep_thm (no_vars ctxt thm)))*}
126
+ − 199
+ − 200
text {*
190
+ − 201
Theorem @{thm [source] conjI} looks now as follows
+ − 202
+ − 203
@{ML_response_fake [display, gray]
+ − 204
"warning (str_of_thm_raw @{context} @{thm conjI})"
+ − 205
"\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q"}
+ − 206
126
+ − 207
Again the function @{ML commas} helps with printing more than one theorem.
+ − 208
*}
+ − 209
+ − 210
ML{*fun str_of_thms ctxt thms =
190
+ − 211
commas (map (str_of_thm ctxt) thms)
+ − 212
+ − 213
fun str_of_thms_raw ctxt thms =
+ − 214
commas (map (str_of_thm_raw ctxt) thms)*}
126
+ − 215
+ − 216
section {* Combinators\label{sec:combinators} *}
+ − 217
+ − 218
text {*
131
+ − 219
For beginners perhaps the most puzzling parts in the existing code of Isabelle are
126
+ − 220
the combinators. At first they seem to greatly obstruct the
+ − 221
comprehension of the code, but after getting familiar with them, they
+ − 222
actually ease the understanding and also the programming.
+ − 223
+ − 224
The simplest combinator is @{ML I}, which is just the identity function defined as
+ − 225
*}
+ − 226
+ − 227
ML{*fun I x = x*}
+ − 228
+ − 229
text {* Another simple combinator is @{ML K}, defined as *}
+ − 230
+ − 231
ML{*fun K x = fn _ => x*}
+ − 232
+ − 233
text {*
+ − 234
@{ML K} ``wraps'' a function around the argument @{text "x"}. However, this
+ − 235
function ignores its argument. As a result, @{ML K} defines a constant function
+ − 236
always returning @{text x}.
+ − 237
+ − 238
The next combinator is reverse application, @{ML "|>"}, defined as:
+ − 239
*}
+ − 240
+ − 241
ML{*fun x |> f = f x*}
+ − 242
+ − 243
text {* While just syntactic sugar for the usual function application,
+ − 244
the purpose of this combinator is to implement functions in a
+ − 245
``waterfall fashion''. Consider for example the function *}
+ − 246
+ − 247
ML %linenosgray{*fun inc_by_five x =
+ − 248
x |> (fn x => x + 1)
+ − 249
|> (fn x => (x, x))
+ − 250
|> fst
+ − 251
|> (fn x => x + 4)*}
+ − 252
+ − 253
text {*
+ − 254
which increments its argument @{text x} by 5. It does this by first incrementing
+ − 255
the argument by 1 (Line 2); then storing the result in a pair (Line 3); taking
+ − 256
the first component of the pair (Line 4) and finally incrementing the first
+ − 257
component by 4 (Line 5). This kind of cascading manipulations of values is quite
+ − 258
common when dealing with theories (for example by adding a definition, followed by
+ − 259
lemmas and so on). The reverse application allows you to read what happens in
+ − 260
a top-down manner. This kind of coding should also be familiar,
149
+ − 261
if you have been exposed to Haskell's do-notation. Writing the function @{ML inc_by_five} using
126
+ − 262
the reverse application is much clearer than writing
+ − 263
*}
+ − 264
+ − 265
ML{*fun inc_by_five x = fst ((fn x => (x, x)) (x + 1)) + 4*}
+ − 266
+ − 267
text {* or *}
+ − 268
+ − 269
ML{*fun inc_by_five x =
+ − 270
((fn x => x + 4) o fst o (fn x => (x, x)) o (fn x => x + 1)) x*}
+ − 271
+ − 272
text {* and typographically more economical than *}
+ − 273
+ − 274
ML{*fun inc_by_five x =
+ − 275
let val y1 = x + 1
+ − 276
val y2 = (y1, y1)
+ − 277
val y3 = fst y2
+ − 278
val y4 = y3 + 4
+ − 279
in y4 end*}
+ − 280
+ − 281
text {*
+ − 282
Another reason why the let-bindings in the code above are better to be
+ − 283
avoided: it is more than easy to get the intermediate values wrong, not to
+ − 284
mention the nightmares the maintenance of this code causes!
+ − 285
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 286
In the context of Isabelle, a ``real world'' example for a function written in
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 287
the waterfall fashion might be the following code:
177
+ − 288
*}
126
+ − 289
177
+ − 290
ML %linenosgray{*fun apply_fresh_args pred ctxt =
+ − 291
pred |> fastype_of
+ − 292
|> binder_types
+ − 293
|> map (pair "z")
+ − 294
|> Variable.variant_frees ctxt [pred]
+ − 295
|> map Free
+ − 296
|> (curry list_comb) pred *}
126
+ − 297
177
+ − 298
text {*
184
+ − 299
This code extracts the argument types of a given
+ − 300
predicate and then generates for each argument type a distinct variable; finally it
183
+ − 301
applies the generated variables to the predicate. For example
+ − 302
+ − 303
@{ML_response_fake [display,gray]
+ − 304
"apply_fresh_args @{term \"P::nat \<Rightarrow> int \<Rightarrow> unit \<Rightarrow> bool\"} @{context}
+ − 305
|> Syntax.string_of_term @{context}
+ − 306
|> warning"
+ − 307
"P z za zb"}
177
+ − 308
184
+ − 309
You can read off this behaviour from how @{ML apply_fresh_args} is
183
+ − 310
coded: in Line 2, the function @{ML fastype_of} calculates the type of the
184
+ − 311
predicate; @{ML binder_types} in the next line produces the list of argument
+ − 312
types (in the case above the list @{text "[nat, int, unit]"}); Line 4
+ − 313
pairs up each type with the string @{text "z"}; the
183
+ − 314
function @{ML variant_frees in Variable} generates for each @{text "z"} a
184
+ − 315
unique name avoiding the given @{text pred}; the list of name-type pairs is turned
+ − 316
into a list of variable terms in Line 6, which in the last line is applied
183
+ − 317
by the function @{ML list_comb} to the predicate. We have to use the
+ − 318
function @{ML curry}, because @{ML list_comb} expects the predicate and the
184
+ − 319
variables list as a pair.
177
+ − 320
+ − 321
The combinator @{ML "#>"} is the reverse function composition. It can be
+ − 322
used to define the following function
126
+ − 323
*}
+ − 324
+ − 325
ML{*val inc_by_six =
+ − 326
(fn x => x + 1)
+ − 327
#> (fn x => x + 2)
+ − 328
#> (fn x => x + 3)*}
+ − 329
+ − 330
text {*
+ − 331
which is the function composed of first the increment-by-one function and then
+ − 332
increment-by-two, followed by increment-by-three. Again, the reverse function
+ − 333
composition allows you to read the code top-down.
+ − 334
+ − 335
The remaining combinators described in this section add convenience for the
+ − 336
``waterfall method'' of writing functions. The combinator @{ML tap} allows
+ − 337
you to get hold of an intermediate result (to do some side-calculations for
+ − 338
instance). The function
+ − 339
+ − 340
*}
+ − 341
+ − 342
ML %linenosgray{*fun inc_by_three x =
+ − 343
x |> (fn x => x + 1)
+ − 344
|> tap (fn x => tracing (makestring x))
+ − 345
|> (fn x => x + 2)*}
+ − 346
+ − 347
text {*
+ − 348
increments the argument first by @{text "1"} and then by @{text "2"}. In the
+ − 349
middle (Line 3), however, it uses @{ML tap} for printing the ``plus-one''
+ − 350
intermediate result inside the tracing buffer. The function @{ML tap} can
+ − 351
only be used for side-calculations, because any value that is computed
+ − 352
cannot be merged back into the ``main waterfall''. To do this, you can use
+ − 353
the next combinator.
+ − 354
+ − 355
The combinator @{ML "`"} is similar to @{ML tap}, but applies a function to the value
+ − 356
and returns the result together with the value (as a pair). For example
+ − 357
the function
+ − 358
*}
+ − 359
+ − 360
ML{*fun inc_as_pair x =
+ − 361
x |> `(fn x => x + 1)
+ − 362
|> (fn (x, y) => (x, y + 1))*}
+ − 363
+ − 364
text {*
+ − 365
takes @{text x} as argument, and then increments @{text x}, but also keeps
+ − 366
@{text x}. The intermediate result is therefore the pair @{ML "(x + 1, x)"
+ − 367
for x}. After that, the function increments the right-hand component of the
+ − 368
pair. So finally the result will be @{ML "(x + 1, x + 1)" for x}.
+ − 369
+ − 370
The combinators @{ML "|>>"} and @{ML "||>"} are defined for
+ − 371
functions manipulating pairs. The first applies the function to
+ − 372
the first component of the pair, defined as
+ − 373
*}
+ − 374
+ − 375
ML{*fun (x, y) |>> f = (f x, y)*}
+ − 376
+ − 377
text {*
+ − 378
and the second combinator to the second component, defined as
+ − 379
*}
+ − 380
+ − 381
ML{*fun (x, y) ||> f = (x, f y)*}
+ − 382
+ − 383
text {*
+ − 384
With the combinator @{ML "|->"} you can re-combine the elements from a pair.
+ − 385
This combinator is defined as
+ − 386
*}
+ − 387
+ − 388
ML{*fun (x, y) |-> f = f x y*}
+ − 389
+ − 390
text {* and can be used to write the following roundabout version
+ − 391
of the @{text double} function:
+ − 392
*}
+ − 393
+ − 394
ML{*fun double x =
+ − 395
x |> (fn x => (x, x))
+ − 396
|-> (fn x => fn y => x + y)*}
+ − 397
+ − 398
text {*
+ − 399
Recall that @{ML "|>"} is the reverse function applications. Recall also that the related
+ − 400
reverse function composition is @{ML "#>"}. In fact all the combinators @{ML "|->"},
+ − 401
@{ML "|>>"} and @{ML "||>"} described above have related combinators for function
+ − 402
composition, namely @{ML "#->"}, @{ML "#>>"} and @{ML "##>"}. Using @{ML "#->"},
+ − 403
for example, the function @{text double} can also be written as:
+ − 404
*}
+ − 405
+ − 406
ML{*val double =
+ − 407
(fn x => (x, x))
+ − 408
#-> (fn x => fn y => x + y)*}
+ − 409
+ − 410
text {*
+ − 411
+ − 412
(FIXME: find a good exercise for combinators)
127
+ − 413
+ − 414
\begin{readmore}
+ − 415
The most frequently used combinator are defined in the files @{ML_file "Pure/library.ML"}
+ − 416
and @{ML_file "Pure/General/basics.ML"}. Also \isccite{sec:ML-linear-trans}
+ − 417
contains further information about combinators.
+ − 418
\end{readmore}
126
+ − 419
15
+ − 420
*}
+ − 421
10
+ − 422
2
+ − 423
section {* Antiquotations *}
+ − 424
+ − 425
text {*
49
+ − 426
The main advantage of embedding all code in a theory is that the code can
58
+ − 427
contain references to entities defined on the logical level of Isabelle. By
+ − 428
this we mean definitions, theorems, terms and so on. This kind of reference is
+ − 429
realised with antiquotations. For example, one can print out the name of the current
49
+ − 430
theory by typing
+ − 431
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 432
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 433
@{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
+ − 434
5
+ − 435
where @{text "@{theory}"} is an antiquotation that is substituted with the
49
+ − 436
current theory (remember that we assumed we are inside the theory
89
+ − 437
@{text FirstSteps}). The name of this theory can be extracted using
49
+ − 438
the function @{ML "Context.theory_name"}.
5
+ − 439
89
+ − 440
Note, however, that antiquotations are statically linked, that is their value is
12
+ − 441
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
+ − 442
*}
5
+ − 443
69
+ − 444
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
+ − 445
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 446
text {*
2
+ − 447
89
+ − 448
does \emph{not} return the name of the current theory, if it is run in a
5
+ − 449
different theory. Instead, the code above defines the constant function
58
+ − 450
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
+ − 451
function is called. Operationally speaking, the antiquotation @{text "@{theory}"} is
5
+ − 452
\emph{not} replaced with code that will look up the current theory in
+ − 453
some data structure and return it. Instead, it is literally
+ − 454
replaced with the value representing the theory name.
2
+ − 455
132
+ − 456
In a similar way you can use antiquotations to refer to proved theorems:
133
+ − 457
@{text "@{thm \<dots>}"} for a single theorem
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 458
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 459
@{ML_response_fake [display,gray] "@{thm allI}" "(\<And>x. ?P x) \<Longrightarrow> \<forall>x. ?P x"}
75
+ − 460
133
+ − 461
and @{text "@{thms \<dots>}"} for more than one
132
+ − 462
+ − 463
@{ML_response_fake [display,gray] "@{thms conj_ac}"
+ − 464
"(?P \<and> ?Q) = (?Q \<and> ?P)
+ − 465
(?P \<and> ?Q \<and> ?R) = (?Q \<and> ?P \<and> ?R)
+ − 466
((?P \<and> ?Q) \<and> ?R) = (?P \<and> ?Q \<and> ?R)"}
+ − 467
149
+ − 468
You can also refer to the current simpset. To illustrate this we implement the
132
+ − 469
function that extracts the theorem names stored in a simpset.
131
+ − 470
*}
75
+ − 471
149
+ − 472
ML{*fun get_thm_names_from_ss simpset =
131
+ − 473
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
+ − 474
val {simps,...} = MetaSimplifier.dest_ss simpset
70
+ − 475
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
+ − 476
map #1 simps
131
+ − 477
end*}
54
+ − 478
131
+ − 479
text {*
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
+ − 480
The function @{ML dest_ss in MetaSimplifier} returns a record containing all
184
+ − 481
information stored in the simpset. We are only interested in the names of the
+ − 482
simp-rules. So now you can feed in the current simpset into this function.
+ − 483
The simpset can be referred to using the antiquotation @{text "@{simpset}"}.
81
+ − 484
131
+ − 485
@{ML_response_fake [display,gray]
149
+ − 486
"get_thm_names_from_ss @{simpset}"
+ − 487
"[\"Nat.of_nat_eq_id\", \"Int.of_int_eq_id\", \"Nat.One_nat_def\", \<dots>]"}
10
+ − 488
156
+ − 489
Again, this way or referencing simpsets makes you independent from additions
+ − 490
of lemmas to the simpset by the user that potentially cause loops.
+ − 491
192
+ − 492
On the ML-level of Isabelle, you often have to work with qualified names;
+ − 493
these are strings with some additional information, such positional information
+ − 494
and qualifiers. Such bindings can be generated with the antiquotation
+ − 495
@{text "@{bindin \<dots>}"}.
+ − 496
+ − 497
@{ML_response [display,gray]
+ − 498
"@{binding \"name\"}"
+ − 499
"name"}
+ − 500
+ − 501
An example where a binding is needed is the function @{ML define in LocalTheory}.
+ − 502
Below this function defines the constant @{term "TrueConj"} as the conjunction
+ − 503
@{term "True \<and> True"}.
+ − 504
*}
+ − 505
+ − 506
local_setup %gray {*
+ − 507
snd o LocalTheory.define Thm.internalK
+ − 508
((@{binding "TrueConj"}, NoSyn),
+ − 509
(Attrib.empty_binding, @{term "True \<and> True"})) *}
+ − 510
+ − 511
text {*
75
+ − 512
While antiquotations have many applications, they were originally introduced in order
89
+ − 513
to avoid explicit bindings for theorems such as:
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 514
*}
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 515
69
+ − 516
ML{*val allI = thm "allI" *}
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 517
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 518
text {*
78
+ − 519
These bindings are difficult to maintain and also can be accidentally
156
+ − 520
overwritten by the user. This often broke Isabelle
49
+ − 521
packages. Antiquotations solve this problem, since they are ``linked''
89
+ − 522
statically at compile-time. However, this static linkage also limits their
+ − 523
usefulness in cases where data needs to be build up dynamically. In the
149
+ − 524
course of this chapter you will learn more about these antiquotations:
122
+ − 525
they can simplify Isabelle programming since one can directly access all
89
+ − 526
kinds of logical elements from th ML-level.
2
+ − 527
*}
+ − 528
15
+ − 529
section {* Terms and Types *}
2
+ − 530
+ − 531
text {*
122
+ − 532
One way to construct terms of Isabelle is by using the antiquotation
89
+ − 533
\mbox{@{text "@{term \<dots>}"}}. For example:
2
+ − 534
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 535
@{ML_response [display,gray]
75
+ − 536
"@{term \"(a::nat) + b = c\"}"
+ − 537
"Const (\"op =\", \<dots>) $
149
+ − 538
(Const (\"HOL.plus_class.plus\", \<dots>) $ \<dots> $ \<dots>) $ \<dots>"}
2
+ − 539
50
+ − 540
This will show the term @{term "(a::nat) + b = c"}, but printed using the internal
13
+ − 541
representation of this term. This internal representation corresponds to the
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 542
datatype @{ML_type "term"}.
2
+ − 543
34
+ − 544
The internal representation of terms uses the usual de Bruijn index mechanism where bound
11
+ − 545
variables are represented by the constructor @{ML Bound}. The index in @{ML Bound} refers to
10
+ − 546
the number of Abstractions (@{ML Abs}) we have to skip until we hit the @{ML Abs} that
12
+ − 547
binds the corresponding variable. However, in Isabelle the names of bound variables are
149
+ − 548
kept at abstractions for printing purposes, and so should be treated only as ``comments''.
123
+ − 549
Application in Isabelle is realised with the term-constructor @{ML $}.
10
+ − 550
2
+ − 551
\begin{readmore}
13
+ − 552
Terms are described in detail in \isccite{sec:terms}. Their
78
+ − 553
definition and many useful operations are implemented in @{ML_file "Pure/term.ML"}.
2
+ − 554
\end{readmore}
+ − 555
13
+ − 556
Sometimes the internal representation of terms can be surprisingly different
157
+ − 557
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
+ − 558
parsing/type-checking/pretty printing can be quite elaborate.
2
+ − 559
10
+ − 560
\begin{exercise}
2
+ − 561
Look at the internal term representation of the following terms, and
89
+ − 562
find out why they are represented like this:
2
+ − 563
+ − 564
\begin{itemize}
+ − 565
\item @{term "case x of 0 \<Rightarrow> 0 | Suc y \<Rightarrow> y"}
+ − 566
\item @{term "\<lambda>(x,y). P y x"}
+ − 567
\item @{term "{ [x::int] | x. x \<le> -2 }"}
+ − 568
\end{itemize}
+ − 569
+ − 570
Hint: The third term is already quite big, and the pretty printer
+ − 571
may omit parts of it by default. If you want to see all of it, you
122
+ − 572
can use the following ML-function to set the printing depth to a higher
+ − 573
value:
12
+ − 574
75
+ − 575
@{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
+ − 576
\end{exercise}
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 577
13
+ − 578
The antiquotation @{text "@{prop \<dots>}"} constructs terms of propositional type,
50
+ − 579
inserting the invisible @{text "Trueprop"}-coercions whenever necessary.
68
+ − 580
Consider for example the pairs
12
+ − 581
126
+ − 582
@{ML_response [display,gray] "(@{term \"P x\"}, @{prop \"P x\"})"
+ − 583
"(Free (\"P\", \<dots>) $ Free (\"x\", \<dots>),
149
+ − 584
Const (\"Trueprop\", \<dots>) $ (Free (\"P\", \<dots>) $ Free (\"x\", \<dots>)))"}
65
+ − 585
108
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 586
where a coercion is inserted in the second component and
12
+ − 587
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 588
@{ML_response [display,gray] "(@{term \"P x \<Longrightarrow> Q x\"}, @{prop \"P x \<Longrightarrow> Q x\"})"
65
+ − 589
"(Const (\"==>\", \<dots>) $ \<dots> $ \<dots>, Const (\"==>\", \<dots>) $ \<dots> $ \<dots>)"}
12
+ − 590
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 591
where it is not (since it is already constructed by a meta-implication).
19
+ − 592
89
+ − 593
Types can be constructed using the antiquotation @{text "@{typ \<dots>}"}. For example:
39
631d12c25bde
substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 594
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 595
@{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
+ − 596
19
+ − 597
\begin{readmore}
+ − 598
Types are described in detail in \isccite{sec:types}. Their
78
+ − 599
definition and many useful operations are implemented
+ − 600
in @{ML_file "Pure/type.ML"}.
19
+ − 601
\end{readmore}
47
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 602
*}
19
+ − 603
+ − 604
156
+ − 605
section {* Constructing Terms and Types Manually\label{sec:terms_types_manually} *}
12
+ − 606
+ − 607
text {*
81
+ − 608
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
+ − 609
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
+ − 610
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
+ − 611
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
+ − 612
@{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
+ − 613
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 614
*}
12
+ − 615
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 616
ML{*fun make_imp P Q =
131
+ − 617
let
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 618
val x = Free ("x", @{typ nat})
131
+ − 619
in
+ − 620
Logic.all x (Logic.mk_implies (P $ x, Q $ x))
+ − 621
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
+ − 622
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 623
text {*
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 624
The reason is that you cannot pass the arguments @{term P} and @{term Q}
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 625
into an antiquotation. For example 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
+ − 626
*}
13
+ − 627
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 628
ML{*fun make_wrong_imp P Q = @{prop "\<And>(x::nat). P x \<Longrightarrow> Q x"} *}
12
+ − 629
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 630
text {*
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 631
To see this apply @{text "@{term S}"} and @{text "@{term T}"}
75
+ − 632
to both functions. With @{ML make_imp} we 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
+ − 633
the given arguments
65
+ − 634
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 635
@{ML_response [display,gray] "make_imp @{term S} @{term T}"
162
+ − 636
"Const \<dots> $
+ − 637
Abs (\"x\", Type (\"nat\",[]),
+ − 638
Const \<dots> $ (Free (\"S\",\<dots>) $ \<dots>) $ (Free (\"T\",\<dots>) $ \<dots>))"}
68
+ − 639
81
+ − 640
whereas with @{ML make_wrong_imp} we obtain a term involving the @{term "P"}
68
+ − 641
and @{text "Q"} from the antiquotation.
+ − 642
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 643
@{ML_response [display,gray] "make_wrong_imp @{term S} @{term T}"
162
+ − 644
"Const \<dots> $
+ − 645
Abs (\"x\", \<dots>,
+ − 646
Const \<dots> $ (Const \<dots> $ (Free (\"P\",\<dots>) $ \<dots>)) $
192
+ − 647
(Const \<dots> $ (Free (\"Q\",\<dots>) $ \<dots>)))"}
65
+ − 648
192
+ − 649
There are a number of handy functions that are frequently used for
+ − 650
constructing terms. One is the function @{ML list_comb} which a term
+ − 651
and a list of terms as argument, and produces as output the term
+ − 652
list applied to the term. For example
+ − 653
+ − 654
@{ML_response_fake [display,gray]
+ − 655
"list_comb (@{term \"P::nat\"}, [@{term \"True\"}, @{term \"False\"}])"
+ − 656
"Free (\"P\", \"nat\") $ Const (\"True\", \"bool\") $ Const (\"False\", \"bool\")"}
+ − 657
+ − 658
(FIXME: handy functions for constructing terms: @{ML lambda},
190
+ − 659
@{ML subst_free})
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 660
*}
49
+ − 661
192
+ − 662
ML {* lambda @{term "x::nat"} @{term "P x"}*}
+ − 663
49
+ − 664
+ − 665
text {*
192
+ − 666
As can be seen this function does not take the typing annotation into account.
+ − 667
13
+ − 668
\begin{readmore}
89
+ − 669
There are many functions in @{ML_file "Pure/term.ML"}, @{ML_file "Pure/logic.ML"} and
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
+ − 670
@{ML_file "HOL/Tools/hologic.ML"} that make such manual constructions of terms
49
+ − 671
and types easier.\end{readmore}
13
+ − 672
+ − 673
Have a look at these files and try to solve the following two exercises:
11
+ − 674
*}
+ − 675
+ − 676
text {*
+ − 677
13
+ − 678
\begin{exercise}\label{fun:revsum}
58
+ − 679
Write a function @{text "rev_sum : term -> term"} that takes a
122
+ − 680
term of the form @{text "t\<^isub>1 + t\<^isub>2 + \<dots> + t\<^isub>n"} (whereby @{text "n"} might be zero)
11
+ − 681
and returns the reversed sum @{text "t\<^isub>n + \<dots> + t\<^isub>2 + t\<^isub>1"}. Assume
+ − 682
the @{text "t\<^isub>i"} can be arbitrary expressions and also note that @{text "+"}
13
+ − 683
associates to the left. Try your function on some examples.
11
+ − 684
\end{exercise}
+ − 685
15
+ − 686
\begin{exercise}\label{fun:makesum}
11
+ − 687
Write a function which takes two terms representing natural numbers
75
+ − 688
in unary notation (like @{term "Suc (Suc (Suc 0))"}), and produce the
11
+ − 689
number representing their sum.
+ − 690
\end{exercise}
+ − 691
122
+ − 692
There are a few subtle issues with constants. They usually crop up when
149
+ − 693
pattern matching terms or types, or when constructing them. While it is perfectly ok
122
+ − 694
to write the function @{text is_true} as follows
+ − 695
*}
+ − 696
+ − 697
ML{*fun is_true @{term True} = true
+ − 698
| is_true _ = false*}
+ − 699
+ − 700
text {*
+ − 701
this does not work for picking out @{text "\<forall>"}-quantified terms. Because
+ − 702
the function
+ − 703
*}
+ − 704
+ − 705
ML{*fun is_all (@{term All} $ _) = true
+ − 706
| is_all _ = false*}
+ − 707
+ − 708
text {*
123
+ − 709
will not correctly match the formula @{prop "\<forall>x::nat. P x"}:
122
+ − 710
+ − 711
@{ML_response [display,gray] "is_all @{term \"\<forall>x::nat. P x\"}" "false"}
+ − 712
+ − 713
The problem is that the @{text "@term"}-antiquotation in the pattern
123
+ − 714
fixes the type of the constant @{term "All"} to be @{typ "('a \<Rightarrow> bool) \<Rightarrow> bool"} for
122
+ − 715
an arbitrary, but fixed type @{typ "'a"}. A properly working alternative
+ − 716
for this function is
+ − 717
*}
+ − 718
+ − 719
ML{*fun is_all (Const ("All", _) $ _) = true
+ − 720
| is_all _ = false*}
+ − 721
+ − 722
text {*
+ − 723
because now
+ − 724
+ − 725
@{ML_response [display,gray] "is_all @{term \"\<forall>x::nat. P x\"}" "true"}
+ − 726
149
+ − 727
matches correctly (the first wildcard in the pattern matches any type and the
+ − 728
second any term).
122
+ − 729
123
+ − 730
However there is still a problem: consider the similar function that
131
+ − 731
attempts to pick out @{text "Nil"}-terms:
122
+ − 732
*}
+ − 733
+ − 734
ML{*fun is_nil (Const ("Nil", _)) = true
+ − 735
| is_nil _ = false *}
+ − 736
+ − 737
text {*
123
+ − 738
Unfortunately, also this function does \emph{not} work as expected, since
122
+ − 739
+ − 740
@{ML_response [display,gray] "is_nil @{term \"Nil\"}" "false"}
+ − 741
123
+ − 742
The problem is that on the ML-level the name of a constant is more
149
+ − 743
subtle than you might expect. The function @{ML is_all} worked correctly,
123
+ − 744
because @{term "All"} is such a fundamental constant, which can be referenced
+ − 745
by @{ML "Const (\"All\", some_type)" for some_type}. However, if you look at
122
+ − 746
+ − 747
@{ML_response [display,gray] "@{term \"Nil\"}" "Const (\"List.list.Nil\", \<dots>)"}
+ − 748
131
+ − 749
the name of the constant @{text "Nil"} depends on the theory in which the
128
+ − 750
term constructor is defined (@{text "List"}) and also in which datatype
+ − 751
(@{text "list"}). Even worse, some constants have a name involving
+ − 752
type-classes. Consider for example the constants for @{term "zero"} and
131
+ − 753
\mbox{@{text "(op *)"}}:
122
+ − 754
+ − 755
@{ML_response [display,gray] "(@{term \"0::nat\"}, @{term \"op *\"})"
126
+ − 756
"(Const (\"HOL.zero_class.zero\", \<dots>),
122
+ − 757
Const (\"HOL.times_class.times\", \<dots>))"}
+ − 758
123
+ − 759
While you could use the complete name, for example
+ − 760
@{ML "Const (\"List.list.Nil\", some_type)" for some_type}, for referring to or
122
+ − 761
matching against @{text "Nil"}, this would make the code rather brittle.
+ − 762
The reason is that the theory and the name of the datatype can easily change.
123
+ − 763
To make the code more robust, it is better to use the antiquotation
122
+ − 764
@{text "@{const_name \<dots>}"}. With this antiquotation you can harness the
123
+ − 765
variable parts of the constant's name. Therefore a functions for
+ − 766
matching against constants that have a polymorphic type should
+ − 767
be written as follows.
122
+ − 768
*}
+ − 769
+ − 770
ML{*fun is_nil_or_all (Const (@{const_name "Nil"}, _)) = true
+ − 771
| is_nil_or_all (Const (@{const_name "All"}, _) $ _) = true
+ − 772
| is_nil_or_all _ = false *}
+ − 773
+ − 774
text {*
131
+ − 775
Occasional you have to calculate what the ``base'' name of a given
124
+ − 776
constant is. For this you can use the function @{ML Sign.extern_const} or
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
+ − 777
@{ML Long_Name.base_name}. For example:
124
+ − 778
+ − 779
@{ML_response [display,gray] "Sign.extern_const @{theory} \"List.list.Nil\"" "\"Nil\""}
122
+ − 780
124
+ − 781
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
+ − 782
the smallest name that is still unique, whereas @{ML base_name in Long_Name} always
124
+ − 783
strips off all qualifiers.
122
+ − 784
+ − 785
\begin{readmore}
162
+ − 786
Functions about naming are implemented in @{ML_file "Pure/General/name_space.ML"};
+ − 787
functions about signatures in @{ML_file "Pure/sign.ML"}.
122
+ − 788
\end{readmore}
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 789
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 790
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
+ − 791
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
+ − 792
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
+ − 793
type is as follows:
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 794
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 795
*}
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 796
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 797
ML{*fun make_fun_type tau1 tau2 = Type ("fun", [tau1, tau2]) *}
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 798
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 799
text {* This can be equally written with the combinator @{ML "-->"} as: *}
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 800
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 801
ML{*fun make_fun_type tau1 tau2 = tau1 --> tau2 *}
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 802
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 803
text {*
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 804
A handy function for manipulating terms is @{ML map_types}: it takes a
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 805
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
+ − 806
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
+ − 807
*}
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 808
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 809
ML{*fun nat_to_int t =
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 810
(case t of
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 811
@{typ nat} => @{typ int}
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 812
| Type (s, ts) => Type (s, map nat_to_int ts)
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 813
| _ => t)*}
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 814
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 815
text {*
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 816
An example as follows:
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 817
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 818
@{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
+ − 819
"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
+ − 820
"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
+ − 821
$ 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
+ − 822
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 823
(FIXME: readmore about types)
122
+ − 824
*}
162
+ − 825
122
+ − 826
124
+ − 827
section {* Type-Checking *}
+ − 828
+ − 829
text {*
+ − 830
131
+ − 831
You can freely construct and manipulate @{ML_type "term"}s and @{ML_type
+ − 832
typ}es, since they are just arbitrary unchecked trees. However, you
+ − 833
eventually want to see if a term is well-formed, or type-checks, relative to
+ − 834
a theory. Type-checking is done via the function @{ML cterm_of}, which
+ − 835
converts a @{ML_type term} into a @{ML_type cterm}, a \emph{certified}
+ − 836
term. Unlike @{ML_type term}s, which are just trees, @{ML_type "cterm"}s are
+ − 837
abstract objects that are guaranteed to be type-correct, and they can only
+ − 838
be constructed via ``official interfaces''.
+ − 839
124
+ − 840
+ − 841
Type-checking is always relative to a theory context. For now we use
+ − 842
the @{ML "@{theory}"} antiquotation to get hold of the current theory.
+ − 843
For example you can write:
+ − 844
149
+ − 845
@{ML_response_fake [display,gray] "cterm_of @{theory} @{term \"(a::nat) + b = c\"}" "a + b = c"}
124
+ − 846
+ − 847
This can also be written with an antiquotation:
+ − 848
+ − 849
@{ML_response_fake [display,gray] "@{cterm \"(a::nat) + b = c\"}" "a + b = c"}
+ − 850
+ − 851
Attempting to obtain the certified term for
+ − 852
+ − 853
@{ML_response_fake_both [display,gray] "@{cterm \"1 + True\"}" "Type unification failed \<dots>"}
+ − 854
+ − 855
yields an error (since the term is not typable). A slightly more elaborate
+ − 856
example that type-checks is:
+ − 857
+ − 858
@{ML_response_fake [display,gray]
+ − 859
"let
+ − 860
val natT = @{typ \"nat\"}
+ − 861
val zero = @{term \"0::nat\"}
+ − 862
in
+ − 863
cterm_of @{theory}
+ − 864
(Const (@{const_name plus}, natT --> natT --> natT) $ zero $ zero)
+ − 865
end" "0 + 0"}
+ − 866
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 867
In Isabelle not just terms need to be certified, but also types. For example,
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 868
you obtain the certified type for the Isablle type @{typ "nat \<Rightarrow> bool"} on
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 869
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
+ − 870
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 871
@{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
+ − 872
"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
+ − 873
"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
+ − 874
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 875
\begin{readmore}
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 876
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
+ − 877
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
+ − 878
\end{readmore}
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 879
124
+ − 880
\begin{exercise}
+ − 881
Check that the function defined in Exercise~\ref{fun:revsum} returns a
+ − 882
result that type-checks.
+ − 883
\end{exercise}
+ − 884
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 885
Remember Isabelle foolows the Church-style typing for terms, i.e.~a term contains
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 886
enough typing information (constants, free variables and abstractions all have typing
124
+ − 887
information) so that it is always clear what the type of a term is.
+ − 888
Given a well-typed term, the function @{ML type_of} returns the
+ − 889
type of a term. Consider for example:
+ − 890
+ − 891
@{ML_response [display,gray]
+ − 892
"type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
+ − 893
+ − 894
To calculate the type, this function traverses the whole term and will
128
+ − 895
detect any typing inconsistency. For examle changing the type of the variable
149
+ − 896
@{term "x"} from @{typ "nat"} to @{typ "int"} will result in the error message:
124
+ − 897
+ − 898
@{ML_response_fake [display,gray]
+ − 899
"type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})"
+ − 900
"*** Exception- TYPE (\"type_of: type mismatch in application\" \<dots>"}
+ − 901
+ − 902
Since the complete traversal might sometimes be too costly and
149
+ − 903
not necessary, there is the function @{ML fastype_of}, which
+ − 904
also returns the type of a term.
124
+ − 905
+ − 906
@{ML_response [display,gray]
+ − 907
"fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
+ − 908
177
+ − 909
However, efficiency is gained on the expense of skipping some tests. You
124
+ − 910
can see this in the following example
+ − 911
+ − 912
@{ML_response [display,gray]
+ − 913
"fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})" "bool"}
+ − 914
149
+ − 915
where no error is detected.
124
+ − 916
+ − 917
Sometimes it is a bit inconvenient to construct a term with
+ − 918
complete typing annotations, especially in cases where the typing
+ − 919
information is redundant. A short-cut is to use the ``place-holder''
+ − 920
type @{ML "dummyT"} and then let type-inference figure out the
+ − 921
complete type. An example is as follows:
+ − 922
+ − 923
@{ML_response_fake [display,gray]
+ − 924
"let
126
+ − 925
val c = Const (@{const_name \"plus\"}, dummyT)
+ − 926
val o = @{term \"1::nat\"}
+ − 927
val v = Free (\"x\", dummyT)
124
+ − 928
in
126
+ − 929
Syntax.check_term @{context} (c $ o $ v)
124
+ − 930
end"
126
+ − 931
"Const (\"HOL.plus_class.plus\", \"nat \<Rightarrow> nat \<Rightarrow> nat\") $
+ − 932
Const (\"HOL.one_class.one\", \"nat\") $ Free (\"x\", \"nat\")"}
124
+ − 933
+ − 934
Instead of giving explicitly the type for the constant @{text "plus"} and the free
149
+ − 935
variable @{text "x"}, the type-inference filles in the missing information.
124
+ − 936
+ − 937
\begin{readmore}
+ − 938
See @{ML_file "Pure/Syntax/syntax.ML"} where more functions about reading,
160
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 939
checking and pretty-printing of terms are defined. Fuctions related to the
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 940
type inference are implemented in @{ML_file "Pure/type.ML"} and
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 941
@{ML_file "Pure/type_infer.ML"}.
124
+ − 942
\end{readmore}
162
+ − 943
+ − 944
(FIXME: say something about sorts)
124
+ − 945
*}
+ − 946
+ − 947
2
+ − 948
section {* Theorems *}
+ − 949
+ − 950
text {*
50
+ − 951
Just like @{ML_type cterm}s, theorems are abstract objects of type @{ML_type thm}
122
+ − 952
that can only be build by going through interfaces. As a consequence, every proof
124
+ − 953
in Isabelle is correct by construction. This follows the tradition of the LCF approach
+ − 954
\cite{GordonMilnerWadsworth79}.
107
+ − 955
2
+ − 956
78
+ − 957
To see theorems in ``action'', let us give a proof on the ML-level for the following
+ − 958
statement:
10
+ − 959
*}
+ − 960
+ − 961
lemma
+ − 962
assumes assm\<^isub>1: "\<And>(x::nat). P x \<Longrightarrow> Q x"
+ − 963
and assm\<^isub>2: "P t"
13
+ − 964
shows "Q t" (*<*)oops(*>*)
10
+ − 965
+ − 966
text {*
185
+ − 967
The corresponding ML-code is as follows:
10
+ − 968
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 969
@{ML_response_fake [display,gray]
42
+ − 970
"let
138
+ − 971
val assm1 = @{cprop \"\<And>(x::nat). P x \<Longrightarrow> Q x\"}
+ − 972
val assm2 = @{cprop \"(P::nat\<Rightarrow>bool) t\"}
10
+ − 973
+ − 974
val Pt_implies_Qt =
+ − 975
assume assm1
138
+ − 976
|> forall_elim @{cterm \"t::nat\"};
10
+ − 977
+ − 978
val Qt = implies_elim Pt_implies_Qt (assume assm2);
+ − 979
in
+ − 980
Qt
+ − 981
|> implies_intr assm2
+ − 982
|> implies_intr assm1
48
+ − 983
end" "\<lbrakk>\<And>x. P x \<Longrightarrow> Q x; P t\<rbrakk> \<Longrightarrow> Q t"}
12
+ − 984
21
+ − 985
This code-snippet constructs the following proof:
+ − 986
+ − 987
\[
+ − 988
\infer[(@{text "\<Longrightarrow>"}$-$intro)]{\vdash @{prop "(\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> P t \<Longrightarrow> Q t"}}
+ − 989
{\infer[(@{text "\<Longrightarrow>"}$-$intro)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"} \vdash @{prop "P t \<Longrightarrow> Q t"}}
+ − 990
{\infer[(@{text "\<Longrightarrow>"}$-$elim)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"}, @{prop "P t"} \vdash @{prop "Q t"}}
+ − 991
{\infer[(@{text "\<And>"}$-$elim)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"} \vdash @{prop "P t \<Longrightarrow> Q t"}}
+ − 992
{\infer[(assume)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"} \vdash @{prop "\<And>x. P x \<Longrightarrow> Q x"}}{}}
+ − 993
&
+ − 994
\infer[(assume)]{@{prop "P t"} \vdash @{prop "P t"}}{}
+ − 995
}
+ − 996
}
+ − 997
}
+ − 998
\]
+ − 999
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
+ − 1000
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
+ − 1001
yet stored in Isabelle's theorem database. So it cannot be referenced later
128
+ − 1002
on. How to store theorems will be explained in Section~\ref{sec:storing}.
21
+ − 1003
13
+ − 1004
\begin{readmore}
50
+ − 1005
For the functions @{text "assume"}, @{text "forall_elim"} etc
13
+ − 1006
see \isccite{sec:thms}. The basic functions for theorems are defined in
+ − 1007
@{ML_file "Pure/thm.ML"}.
+ − 1008
\end{readmore}
12
+ − 1009
127
+ − 1010
(FIXME: how to add case-names to goal states - maybe in the next section)
10
+ − 1011
*}
+ − 1012
123
+ − 1013
section {* Theorem Attributes *}
+ − 1014
127
+ − 1015
text {*
133
+ − 1016
Theorem attributes are @{text "[simp]"}, @{text "[OF \<dots>]"}, @{text
136
+ − 1017
"[symmetric]"} and so on. Such attributes are \emph{neither} tags \emph{nor} flags
+ − 1018
annotated to theorems, but functions that do further processing once a
+ − 1019
theorem is proven. In particular, it is not possible to find out
+ − 1020
what are all theorems that have a given attribute in common, unless of course
133
+ − 1021
the function behind the attribute stores the theorems in a retrivable
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
+ − 1022
datastructure.
127
+ − 1023
+ − 1024
If you want to print out all currently known attributes a theorem
+ − 1025
can have, you can use the function:
+ − 1026
+ − 1027
@{ML_response_fake [display,gray] "Attrib.print_attributes @{theory}"
+ − 1028
"COMP: direct composition with rules (no lifting)
+ − 1029
HOL.dest: declaration of Classical destruction rule
+ − 1030
HOL.elim: declaration of Classical elimination rule
+ − 1031
\<dots>"}
+ − 1032
136
+ − 1033
To explain how to write your own attribute, let us start with an extremely simple
+ − 1034
version of the attribute @{text "[symmetric]"}. The purpose of this attribute is
+ − 1035
to produce the ``symmetric'' version of an equation. The main function behind
+ − 1036
this attribute is
127
+ − 1037
*}
+ − 1038
133
+ − 1039
ML{*val my_symmetric = Thm.rule_attribute (fn _ => fn thm => thm RS @{thm sym})*}
+ − 1040
+ − 1041
text {*
136
+ − 1042
where the function @{ML "Thm.rule_attribute"} expects a function taking a
149
+ − 1043
context (which we ignore in the code above) and a theorem (@{text thm}), and
+ − 1044
returns another theorem (namely @{text thm} resolved with the lemma
+ − 1045
@{thm [source] sym}: @{thm sym[no_vars]}). The function @{ML "Thm.rule_attribute"} then returns
156
+ − 1046
an attribute.
136
+ − 1047
+ − 1048
Before we can use the attribute, we need to set it up. This can be done
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1049
using the function @{ML Attrib.setup} as follows.
133
+ − 1050
*}
+ − 1051
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1052
setup %gray {* Attrib.setup @{binding "my_sym"}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1053
(Scan.succeed my_symmetric) "applying the sym rule"*}
133
+ − 1054
136
+ − 1055
text {*
149
+ − 1056
The attribute does not expect any further arguments (unlike @{text "[OF
+ − 1057
\<dots>]"}, for example, which can take a list of theorems as argument). Therefore
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1058
we use the parser @{ML Scan.succeed}. Later on we will also consider
149
+ − 1059
attributes taking further arguments. An example for the attribute @{text
+ − 1060
"[my_sym]"} is the proof
136
+ − 1061
*}
+ − 1062
+ − 1063
lemma test[my_sym]: "2 = Suc (Suc 0)" by simp
133
+ − 1064
+ − 1065
text {*
136
+ − 1066
which stores the theorem @{thm test} under the name @{thm [source] test}. We
+ − 1067
can also use the attribute when referring to this theorem.
+ − 1068
+ − 1069
\begin{isabelle}
+ − 1070
\isacommand{thm}~@{text "test[my_sym]"}\\
+ − 1071
@{text "> "}~@{thm test[my_sym]}
+ − 1072
\end{isabelle}
+ − 1073
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
+ − 1074
The purpose of @{ML Thm.rule_attribute} is to directly manipulate theorems.
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
+ − 1075
Another usage of attributes is to add and delete theorems from stored data.
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
+ − 1076
For example the attribute @{text "[simp]"} adds or deletes a theorem from the
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
+ − 1077
current simpset. For these applications, you can use @{ML Thm.declaration_attribute}.
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
+ − 1078
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
+ − 1079
of theorems.
133
+ − 1080
*}
+ − 1081
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
+ − 1082
ML{*val my_thms = ref ([]: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
+ − 1083
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
+ − 1084
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
+ − 1085
A word of warning: such references must not be used in any code that
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
+ − 1086
is meant to be more than just for testing purposes! Here it is only used
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
+ − 1087
to illustrate matters. We will show later how to store data properly without
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
+ − 1088
using references. The function @{ML Thm.declaration_attribute} expects us to
153
+ − 1089
provide two functions that add and delete theorems from this list.
+ − 1090
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
+ − 1091
*}
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
+ − 1092
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
+ − 1093
ML{*fun my_thms_add thm 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
+ − 1094
(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
+ − 1095
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
+ − 1096
fun my_thms_del thm 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
+ − 1097
(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
+ − 1098
133
+ − 1099
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
+ − 1100
These functions take a theorem and a context and, for what we are explaining
156
+ − 1101
here it is sufficient that they just return the context unchanged. They change
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
+ − 1102
however the reference @{ML my_thms}, whereby the function @{ML Thm.add_thm}
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
+ − 1103
adds a theorem if it is not already included in the list, and @{ML
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
+ − 1104
Thm.del_thm} deletes one. Both functions use the predicate @{ML
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
+ − 1105
Thm.eq_thm_prop} which compares theorems according to their proved
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
+ − 1106
propositions (modulo alpha-equivalence).
133
+ − 1107
+ − 1108
156
+ − 1109
You can turn both functions into attributes using
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
+ − 1110
*}
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
+ − 1111
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
+ − 1112
ML{*val my_add = Thm.declaration_attribute my_thms_add
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
+ − 1113
val my_del = Thm.declaration_attribute 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
+ − 1114
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
+ − 1115
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
+ − 1116
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
+ − 1117
*}
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
+ − 1118
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1119
setup %gray {* Attrib.setup @{binding "my_thms"}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1120
(Attrib.add_del my_add my_del) "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
+ − 1121
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
+ − 1122
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
+ − 1123
Now if you prove the lemma attaching the attribute @{text "[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
+ − 1124
*}
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
+ − 1125
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
+ − 1126
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
+ − 1127
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
+ − 1128
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
+ − 1129
then you can see the lemma is added to the initially empty 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
+ − 1130
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
+ − 1131
@{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
+ − 1132
"!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
+ − 1133
160
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1134
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
+ − 1135
*}
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
+ − 1136
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
+ − 1137
declare test[my_thms] trueI_2[my_thms add]
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
+ − 1138
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
+ − 1139
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
+ − 1140
The @{text "add"} is the default operation and does not need
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
+ − 1141
to be given. This declaration will cause the theorem list to be
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
+ − 1142
updated 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
+ − 1143
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
+ − 1144
@{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
+ − 1145
"!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
+ − 1146
"[\"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
+ − 1147
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
+ − 1148
The theorem @{thm [source] trueI_2} only appears once, since the
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
+ − 1149
function @{ML Thm.add_thm} tests for duplicates, before extending
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
+ − 1150
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
+ − 1151
*}
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
+ − 1152
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
+ − 1153
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
+ − 1154
156
+ − 1155
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
+ − 1156
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
+ − 1157
@{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
+ − 1158
"!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
+ − 1159
"[\"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
+ − 1160
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
+ − 1161
We used in this example two functions declared as @{ML Thm.declaration_attribute},
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
+ − 1162
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
+ − 1163
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
+ − 1164
156
+ − 1165
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
+ − 1166
\emph{not} the received way of doing such things. The received way is to
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
+ − 1167
start a ``data slot'' in a context by using the functor @{text GenericDataFun}:
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
+ − 1168
*}
133
+ − 1169
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
+ − 1170
ML {*structure Data = GenericDataFun
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
+ − 1171
(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
+ − 1172
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
+ − 1173
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
+ − 1174
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
+ − 1175
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
+ − 1176
text {*
156
+ − 1177
To use this data slot, you only have to change @{ML my_thms_add} and
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
+ − 1178
@{ML my_thms_del} to:
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
+ − 1179
*}
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
+ − 1180
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
+ − 1181
ML{*val thm_add = Data.map o Thm.add_thm
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
+ − 1182
val thm_del = Data.map o Thm.del_thm*}
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
+ − 1183
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
+ − 1184
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
+ − 1185
where @{ML Data.map} updates the data appropriately in the context. Since storing
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
+ − 1186
theorems in a special list is such a common task, there is the functor @{text NamedThmsFun},
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
+ − 1187
which does most of the work for you. To obtain such a named theorem lists, you just
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
+ − 1188
declare
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
+ − 1189
*}
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
+ − 1190
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
+ − 1191
ML{*structure FooRules = NamedThmsFun
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
+ − 1192
(val name = "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
+ − 1193
val description = "Rules for 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
+ − 1194
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
+ − 1195
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
+ − 1196
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
+ − 1197
*}
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
+ − 1198
177
+ − 1199
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
+ − 1200
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
+ − 1201
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
+ − 1202
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
+ − 1203
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
+ − 1204
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
+ − 1205
modify the theorems.
133
+ − 1206
157
+ − 1207
Furthermore, the facts are made available on the user-level under the dynamic
156
+ − 1208
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
+ − 1209
@{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
+ − 1210
*}
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
+ − 1211
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
+ − 1212
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
+ − 1213
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
+ − 1214
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
+ − 1215
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
+ − 1216
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
+ − 1217
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
+ − 1218
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
+ − 1219
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
+ − 1220
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
+ − 1221
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
+ − 1222
\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
+ − 1223
\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
+ − 1224
@{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
+ − 1225
@{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
+ − 1226
\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
+ − 1227
156
+ − 1228
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
+ − 1229
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
+ − 1230
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
+ − 1231
@{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
+ − 1232
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
+ − 1233
\begin{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
+ − 1234
For more information see @{ML_file "Pure/Tools/named_thms.ML"} and also
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
+ − 1235
the recipe in Section~\ref{recipe:storingdata} about storing arbitrary
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
+ − 1236
data.
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
+ − 1237
\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
+ − 1238
156
+ − 1239
(FIXME What are: @{text "theory_attributes"}, @{text "proof_attributes"}?)
133
+ − 1240
+ − 1241
+ − 1242
\begin{readmore}
+ − 1243
FIXME: @{ML_file "Pure/more_thm.ML"} @{ML_file "Pure/Isar/attrib.ML"}
+ − 1244
\end{readmore}
+ − 1245
*}
+ − 1246
127
+ − 1247
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1248
section {* Setups (TBD) *}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1249
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1250
text {*
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1251
In the previous section we used \isacommand{setup} in order to make
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1252
a theorem attribute known to Isabelle. What happens behind the scenes
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1253
is that \isacommand{setup} expects a functions of type
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1254
@{ML_type "theory -> theory"}: the input theory is the current theory and the
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1255
output the theory where the theory attribute has been stored.
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1256
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1257
This is a fundamental principle in Isabelle. A similar situation occurs
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1258
for example with declaring constants. The function that declares a
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1259
constant on the ML-level is @{ML Sign.add_consts_i}.
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1260
If you write\footnote{Recall that ML-code needs to be
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1261
enclosed in \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.}
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1262
*}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1263
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1264
ML{*Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] @{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
+ − 1265
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1266
text {*
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1267
for declaring the constant @{text "BAR"} with type @{typ nat} and
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1268
run the code, then you indeed obtain a theory as result. But if you
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1269
query the constant on the Isabelle level using the command \isacommand{term}
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1270
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1271
\begin{isabelle}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1272
\isacommand{term}~@{text [quotes] "BAR"}\\
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1273
@{text "> \"BAR\" :: \"'a\""}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1274
\end{isabelle}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1275
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1276
you do not obtain a constant of type @{typ nat}, but a free variable (printed in
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1277
blue) of polymorphic type. The problem is that the ML-expression above did
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1278
not register the declaration with the current theory. This is what the command
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1279
\isacommand{setup} is for. The constant is properly declared with
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1280
*}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1281
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1282
setup %gray {* Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] *}
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1283
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1284
text {*
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1285
Now
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1286
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1287
\begin{isabelle}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1288
\isacommand{term}~@{text [quotes] "BAR"}\\
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1289
@{text "> \"BAR\" :: \"nat\""}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1290
\end{isabelle}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1291
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1292
returns a (black) constant with the type @{typ nat}.
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1293
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1294
A similar command is \isacommand{local\_setup}, which expects a function
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1295
of type @{ML_type "local_theory -> local_theory"}. Later on we will also
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1296
use the commands \isacommand{method\_setup} for installing methods in the
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1297
current theory and \isacommand{simproc\_setup} for adding new simprocs to
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1298
the current simpset.
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1299
*}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1300
153
+ − 1301
section {* Theories, Contexts and Local Theories (TBD) *}
123
+ − 1302
126
+ − 1303
text {*
+ − 1304
There are theories, proof contexts and local theories (in this order, if you
+ − 1305
want to order them).
+ − 1306
+ − 1307
In contrast to an ordinary theory, which simply consists of a type
+ − 1308
signature, as well as tables for constants, axioms and theorems, a local
+ − 1309
theory also contains additional context information, such as locally fixed
+ − 1310
variables and local assumptions that may be used by the package. The type
+ − 1311
@{ML_type local_theory} is identical to the type of \emph{proof contexts}
+ − 1312
@{ML_type "Proof.context"}, although not every proof context constitutes a
+ − 1313
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
+ − 1314
*}
126
+ − 1315
153
+ − 1316
section {* Storing Theorems\label{sec:storing} (TBD) *}
123
+ − 1317
+ − 1318
text {* @{ML PureThy.add_thms_dynamic} *}
+ − 1319
100
+ − 1320
75
+ − 1321
126
+ − 1322
(* FIXME: some code below *)
89
+ − 1323
+ − 1324
(*<*)
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
+ − 1325
(*
89
+ − 1326
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
+ − 1327
Sign.add_consts_i [(Binding"bar", @{typ "nat"},NoSyn)]
89
+ − 1328
*}
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
+ − 1329
*)
89
+ − 1330
lemma "bar = (1::nat)"
+ − 1331
oops
+ − 1332
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
+ − 1333
(*
89
+ − 1334
setup {*
+ − 1335
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
+ − 1336
#> PureThy.add_defs false [((@{binding "foo_def"},
89
+ − 1337
Logic.mk_equals (Const ("FirstSteps.foo", @{typ "nat"}), @{term "1::nat"})), [])]
+ − 1338
#> snd
+ − 1339
*}
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
+ − 1340
*)
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1341
(*
89
+ − 1342
lemma "foo = (1::nat)"
+ − 1343
apply(simp add: foo_def)
+ − 1344
done
+ − 1345
+ − 1346
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
+ − 1347
*)
89
+ − 1348
(*>*)
+ − 1349
153
+ − 1350
section {* Pretty-Printing (TBD) *}
+ − 1351
+ − 1352
text {*
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
+ − 1353
@{ML Pretty.big_list},
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1354
@{ML Pretty.brk},
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1355
@{ML Pretty.block},
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1356
@{ML Pretty.chunks}
153
+ − 1357
*}
+ − 1358
+ − 1359
section {* Misc (TBD) *}
92
+ − 1360
+ − 1361
ML {*DatatypePackage.get_datatype @{theory} "List.list"*}
+ − 1362
2
+ − 1363
end