32
+ − 1
theory Ind_Interface
129
+ − 2
imports "../Base" "../Parsing" Ind_Prelims Simple_Inductive_Package
32
+ − 3
begin
+ − 4
215
+ − 5
section {* Parsing and Typing the Specification\label{sec:interface} *}
124
+ − 6
+ − 7
text {*
116
+ − 8
To be able to write down the specification in Isabelle, we have to introduce
129
+ − 9
a new command (see Section~\ref{sec:newcommand}). As the keyword for the
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 10
new command we chose \simpleinductive{}. The ``infrastructure'' already
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 11
provides an ``advanced'' feature for this command. For example we will
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 12
be able to declare the locale
127
+ − 13
*}
+ − 14
+ − 15
locale rel =
+ − 16
fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+ − 17
+ − 18
text {*
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 19
and then define the transitive closure and the accessible part of this
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 20
locale as follows:
116
+ − 21
*}
+ − 22
129
+ − 23
simple_inductive (in rel)
+ − 24
trcl'
116
+ − 25
where
127
+ − 26
base: "trcl' x x"
+ − 27
| step: "trcl' x y \<Longrightarrow> R y z \<Longrightarrow> trcl' x z"
116
+ − 28
129
+ − 29
simple_inductive (in rel)
+ − 30
accpart'
116
+ − 31
where
127
+ − 32
accpartI: "(\<And>y. R y x \<Longrightarrow> accpart' y) \<Longrightarrow> accpart' x"
+ − 33
116
+ − 34
text {*
129
+ − 35
\begin{figure}[t]
120
+ − 36
\begin{isabelle}
116
+ − 37
\railnontermfont{\rmfamily\itshape}
+ − 38
\railterm{simpleinductive,where,for}
+ − 39
\railalias{simpleinductive}{\simpleinductive{}}
+ − 40
\railalias{where}{\isacommand{where}}
+ − 41
\railalias{for}{\isacommand{for}}
+ − 42
\begin{rail}
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 43
simpleinductive target?\\ fixes
116
+ − 44
(where (thmdecl? prop + '|'))?
+ − 45
;
+ − 46
\end{rail}
120
+ − 47
\end{isabelle}
+ − 48
\caption{A railroad diagram describing the syntax of \simpleinductive{}.
+ − 49
The \emph{target} indicates an optional locale; the \emph{fixes} are an
+ − 50
\isacommand{and}-separated list of names for the inductive predicates (they
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 51
can also contain typing- and syntax anotations); \emph{prop} stands for a
120
+ − 52
introduction rule with an optional theorem declaration (\emph{thmdecl}).
+ − 53
\label{fig:railroad}}
+ − 54
\end{figure}
116
+ − 55
*}
+ − 56
+ − 57
text {*
129
+ − 58
This leads directly to the railroad diagram shown in
+ − 59
Figure~\ref{fig:railroad} for the syntax of \simpleinductive{}. This diagram
120
+ − 60
more or less translates directly into the parser:
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 61
*}
120
+ − 62
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 63
ML %linenosgray{*val spec_parser =
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 64
OuterParse.fixes --
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 65
Scan.optional
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 66
(OuterParse.$$$ "where" |--
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 67
OuterParse.!!!
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 68
(OuterParse.enum1 "|"
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 69
(SpecParse.opt_thm_name ":" -- OuterParse.prop))) []*}
118
+ − 70
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 71
text {*
120
+ − 72
which we described in Section~\ref{sec:parsingspecs}. If we feed into the
+ − 73
parser the string (which corresponds to our definition of @{term even} and
+ − 74
@{term odd}):
+ − 75
+ − 76
@{ML_response [display,gray]
+ − 77
"let
+ − 78
val input = filtered_input
+ − 79
(\"even and odd \" ^
+ − 80
\"where \" ^
+ − 81
\" even0[intro]: \\\"even 0\\\" \" ^
+ − 82
\"| evenS[intro]: \\\"odd n \<Longrightarrow> even (Suc n)\\\" \" ^
+ − 83
\"| oddS[intro]: \\\"even n \<Longrightarrow> odd (Suc n)\\\"\")
+ − 84
in
+ − 85
parse spec_parser input
+ − 86
end"
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 87
"(([(even, NONE, NoSyn), (odd, NONE, NoSyn)],
120
+ − 88
[((even0,\<dots>), \"\\^E\\^Ftoken\\^Eeven 0\\^E\\^F\\^E\"),
+ − 89
((evenS,\<dots>), \"\\^E\\^Ftoken\\^Eodd n \<Longrightarrow> even (Suc n)\\^E\\^F\\^E\"),
+ − 90
((oddS,\<dots>), \"\\^E\\^Ftoken\\^Eeven n \<Longrightarrow> odd (Suc n)\\^E\\^F\\^E\")]), [])"}
+ − 91
*}
+ − 92
+ − 93
+ − 94
text {*
+ − 95
then we get back a locale (in this case @{ML NONE}), the predicates (with type
+ − 96
and syntax annotations), the parameters (similar as the predicates) and
+ − 97
the specifications of the introduction rules.
+ − 98
+ − 99
+ − 100
+ − 101
This is all the information we
+ − 102
need for calling the package and setting up the keyword. The latter is
+ − 103
done in Lines 6 and 7 in the code below.
116
+ − 104
118
+ − 105
@{ML_chunk [display,gray,linenos] syntax}
120
+ − 106
+ − 107
We call @{ML OuterSyntax.command} with the kind-indicator @{ML
+ − 108
OuterKeyword.thy_decl} since the package does not need to open up any goal
+ − 109
state (see Section~\ref{sec:newcommand}). Note that the predicates and
+ − 110
parameters are at the moment only some ``naked'' variables: they have no
+ − 111
type yet (even if we annotate them with types) and they are also no defined
+ − 112
constants yet (which the predicates will eventually be). In Lines 1 to 4 we
+ − 113
gather the information from the parser to be processed further. The locale
+ − 114
is passed as argument to the function @{ML
+ − 115
Toplevel.local_theory}.\footnote{FIXME Is this already described?} The other
+ − 116
arguments, i.e.~the predicates, parameters and intro rule specifications,
+ − 117
are passed to the function @{ML add_inductive in SimpleInductivePackage}
+ − 118
(Line 4).
116
+ − 119
120
+ − 120
We now come to the second subtask of the package, namely transforming the
+ − 121
parser output into some internal datastructures that can be processed further.
+ − 122
Remember that at the moment the introduction rules are just strings, and even
+ − 123
if the predicates and parameters can contain some typing annotations, they
+ − 124
are not yet in any way reflected in the introduction rules. So the task of
+ − 125
@{ML add_inductive in SimpleInductivePackage} is to transform the strings
+ − 126
into properly typed terms. For this it can use the function
183
+ − 127
@{ML read_spec in Specification}. This function takes some constants
120
+ − 128
with possible typing annotations and some rule specifications and attempts to
+ − 129
find a type according to the given type constraints and the type constraints
+ − 130
by the surrounding (local theory). However this function is a bit
+ − 131
too general for our purposes: we want that each introduction rule has only
+ − 132
name (for example @{text even0} or @{text evenS}), if a name is given at all.
183
+ − 133
The function @{ML read_spec in Specification} however allows more
120
+ − 134
than one rule. Since it is quite convenient to rely on this function (instead of
+ − 135
building your own) we just quick ly write a wrapper function that translates
+ − 136
between our specific format and the general format expected by
183
+ − 137
@{ML read_spec in Specification}. The code of this wrapper is as follows:
120
+ − 138
+ − 139
@{ML_chunk [display,gray,linenos] read_specification}
+ − 140
+ − 141
It takes a list of constants, a list of rule specifications and a local theory
+ − 142
as input. Does the transformation of the rule specifications in Line 3; calls
+ − 143
the function and transforms the now typed rule specifications back into our
+ − 144
format and returns the type parameter and typed rule specifications.
+ − 145
+ − 146
+ − 147
@{ML_chunk [display,gray,linenos] add_inductive}
+ − 148
53
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 149
88
+ − 150
In order to add a new inductive predicate to a theory with the help of our
+ − 151
package, the user must \emph{invoke} it. For every package, there are
+ − 152
essentially two different ways of invoking it, which we will refer to as
+ − 153
\emph{external} and \emph{internal}. By external invocation we mean that the
116
+ − 154
package is called from within a theory document. In this case, the
+ − 155
specification of the inductive predicate, including type annotations and
+ − 156
introduction rules, are given as strings by the user. Before the package can
+ − 157
actually make the definition, the type and introduction rules have to be
+ − 158
parsed. In contrast, internal invocation means that the package is called by
+ − 159
some other package. For example, the function definition package
120
+ − 160
calls the inductive definition package to define the
116
+ − 161
graph of the function. However, it is not a good idea for the function
+ − 162
definition package to pass the introduction rules for the function graph to
+ − 163
the inductive definition package as strings. In this case, it is better to
+ − 164
directly pass the rules to the package as a list of terms, which is more
+ − 165
robust than handling strings that are lacking the additional structure of
+ − 166
terms. These two ways of invoking the package are reflected in its ML
+ − 167
programming interface, which consists of two functions:
+ − 168
88
+ − 169
113
+ − 170
@{ML_chunk [display,gray] SIMPLE_INDUCTIVE_PACKAGE}
53
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 171
*}
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 172
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 173
text {*
124
+ − 174
(FIXME: explain Binding.binding; Attrib.binding somewhere else)
116
+ − 175
+ − 176
88
+ − 177
The function for external invocation of the package is called @{ML
+ − 178
add_inductive in SimpleInductivePackage}, whereas the one for internal
+ − 179
invocation is called @{ML add_inductive_i in SimpleInductivePackage}. Both
+ − 180
of these functions take as arguments the names and types of the inductive
+ − 181
predicates, the names and types of their parameters, the actual introduction
+ − 182
rules and a \emph{local theory}. They return a local theory containing the
116
+ − 183
definition and the induction principle as well introduction rules.
+ − 184
+ − 185
Note that @{ML add_inductive_i in SimpleInductivePackage} expects
88
+ − 186
the types of the predicates and parameters to be specified using the
+ − 187
datatype @{ML_type typ} of Isabelle's logical framework, whereas @{ML
+ − 188
add_inductive in SimpleInductivePackage} expects them to be given as
+ − 189
optional strings. If no string is given for a particular predicate or
+ − 190
parameter, this means that the type should be inferred by the
116
+ − 191
package.
+ − 192
+ − 193
+ − 194
Additional \emph{mixfix syntax} may be associated with the
88
+ − 195
predicates and parameters as well. Note that @{ML add_inductive_i in
+ − 196
SimpleInductivePackage} does not allow mixfix syntax to be associated with
116
+ − 197
parameters, since it can only be used for parsing.\footnote{FIXME: why ist it there then?}
+ − 198
The names of the
88
+ − 199
predicates, parameters and rules are represented by the type @{ML_type
+ − 200
Binding.binding}. Strings can be turned into elements of the type @{ML_type
+ − 201
Binding.binding} using the function @{ML [display] "Binding.name : string ->
+ − 202
Binding.binding"} Each introduction rule is given as a tuple containing its
+ − 203
name, a list of \emph{attributes} and a logical formula. Note that the type
+ − 204
@{ML_type Attrib.binding} used in the list of introduction rules is just a
+ − 205
shorthand for the type @{ML_type "Binding.binding * Attrib.src list"}. The
+ − 206
function @{ML add_inductive_i in SimpleInductivePackage} expects the formula
+ − 207
to be specified using the datatype @{ML_type term}, whereas @{ML
+ − 208
add_inductive in SimpleInductivePackage} expects it to be given as a string.
+ − 209
An attribute specifies additional actions and transformations that should be
+ − 210
applied to a theorem, such as storing it in the rule databases used by
+ − 211
automatic tactics like the simplifier. The code of the package, which will
+ − 212
be described in the following section, will mostly treat attributes as a
+ − 213
black box and just forward them to other functions for storing theorems in
+ − 214
local theories. The implementation of the function @{ML add_inductive in
+ − 215
SimpleInductivePackage} for external invocation of the package is quite
+ − 216
simple. Essentially, it just parses the introduction rules and then passes
+ − 217
them on to @{ML add_inductive_i in SimpleInductivePackage}:
+ − 218
+ − 219
@{ML_chunk [display] add_inductive}
+ − 220
+ − 221
For parsing and type checking the introduction rules, we use the function
+ − 222
+ − 223
@{ML [display] "Specification.read_specification:
76
+ − 224
(Binding.binding * string option * mixfix) list -> (*{variables}*)
176
+ − 225
(Attrib.binding * string list) list -> (*{rules}*)
32
+ − 226
local_theory ->
76
+ − 227
(((Binding.binding * typ) * mixfix) list *
32
+ − 228
(Attrib.binding * term list) list) *
+ − 229
local_theory"}
53
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 230
*}
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 231
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 232
text {*
88
+ − 233
During parsing, both predicates and parameters are treated as variables, so
+ − 234
the lists \verb!preds_syn! and \verb!params_syn! are just appended
183
+ − 235
before being passed to @{ML read_spec in Specification}. Note that the format
+ − 236
for rules supported by @{ML read_spec in Specification} is more general than
88
+ − 237
what is required for our package. It allows several rules to be associated
+ − 238
with one name, and the list of rules can be partitioned into several
+ − 239
sublists. In order for the list \verb!intro_srcs! of introduction rules
183
+ − 240
to be acceptable as an input for @{ML read_spec in Specification}, we first
88
+ − 241
have to turn it into a list of singleton lists. This transformation
+ − 242
has to be reversed later on by applying the function
+ − 243
@{ML [display] "the_single: 'a list -> 'a"}
+ − 244
to the list \verb!specs! containing the parsed introduction rules.
183
+ − 245
The function @{ML read_spec in Specification} also returns the list \verb!vars!
88
+ − 246
of predicates and parameters that contains the inferred types as well.
+ − 247
This list has to be chopped into the two lists \verb!preds_syn'! and
+ − 248
\verb!params_syn'! for predicates and parameters, respectively.
+ − 249
All variables occurring in a rule but not in the list of variables passed to
183
+ − 250
@{ML read_spec in Specification} will be bound by a meta-level universal
88
+ − 251
quantifier.
53
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 252
*}
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 253
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 254
text {*
88
+ − 255
Finally, @{ML read_specification in Specification} also returns another local theory,
+ − 256
but we can safely discard it. As an example, let us look at how we can use this
+ − 257
function to parse the introduction rules of the @{text trcl} predicate:
+ − 258
+ − 259
@{ML_response [display]
32
+ − 260
"Specification.read_specification
55
+ − 261
[(Binding.name \"trcl\", NONE, NoSyn),
+ − 262
(Binding.name \"r\", SOME \"'a \<Rightarrow> 'a \<Rightarrow> bool\", NoSyn)]
176
+ − 263
[((Binding.name \"base\", []), [\"trcl r x x\"]),
+ − 264
((Binding.name \"step\", []), [\"trcl r x y \<Longrightarrow> r y z \<Longrightarrow> trcl r x z\"])]
32
+ − 265
@{context}"
+ − 266
"((\<dots>,
+ − 267
[(\<dots>,
+ − 268
[Const (\"all\", \<dots>) $ Abs (\"x\", TFree (\"'a\", \<dots>),
+ − 269
Const (\"Trueprop\", \<dots>) $
+ − 270
(Free (\"trcl\", \<dots>) $ Free (\"r\", \<dots>) $ Bound 0 $ Bound 0))]),
+ − 271
(\<dots>,
+ − 272
[Const (\"all\", \<dots>) $ Abs (\"x\", TFree (\"'a\", \<dots>),
+ − 273
Const (\"all\", \<dots>) $ Abs (\"y\", TFree (\"'a\", \<dots>),
+ − 274
Const (\"all\", \<dots>) $ Abs (\"z\", TFree (\"'a\", \<dots>),
+ − 275
Const (\"==>\", \<dots>) $
+ − 276
(Const (\"Trueprop\", \<dots>) $
42
+ − 277
(Free (\"trcl\", \<dots>) $ Free (\"r\", \<dots>) $ Bound 2 $ Bound 1)) $
32
+ − 278
(Const (\"==>\", \<dots>) $ \<dots> $ \<dots>))))])]),
+ − 279
\<dots>)
76
+ − 280
: (((Binding.binding * typ) * mixfix) list *
32
+ − 281
(Attrib.binding * term list) list) * local_theory"}
88
+ − 282
+ − 283
In the list of variables passed to @{ML read_specification in Specification}, we have
+ − 284
used the mixfix annotation @{ML NoSyn} to indicate that we do not want to associate any
+ − 285
mixfix syntax with the variable. Moreover, we have only specified the type of \texttt{r},
+ − 286
whereas the type of \texttt{trcl} is computed using type inference.
+ − 287
The local variables \texttt{x}, \texttt{y} and \texttt{z} of the introduction rules
+ − 288
are turned into bound variables with the de Bruijn indices,
+ − 289
whereas \texttt{trcl} and \texttt{r} remain free variables.
+ − 290
53
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 291
*}
32
+ − 292
53
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 293
text {*
88
+ − 294
+ − 295
\paragraph{Parsers for theory syntax}
32
+ − 296
88
+ − 297
Although the function @{ML add_inductive in SimpleInductivePackage} parses terms and types, it still
+ − 298
cannot be used to invoke the package directly from within a theory document.
+ − 299
In order to do this, we have to write another parser. Before we describe
+ − 300
the process of writing parsers for theory syntax in more detail, we first
+ − 301
show some examples of how we would like to use the inductive definition
+ − 302
package.
32
+ − 303
88
+ − 304
+ − 305
The definition of the transitive closure should look as follows:
32
+ − 306
*}
+ − 307
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 308
ML {* SpecParse.opt_thm_name *}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 309
32
+ − 310
text {*
+ − 311
88
+ − 312
A proposition can be parsed using the function @{ML prop in OuterParse}.
+ − 313
Essentially, a proposition is just a string or an identifier, but using the
+ − 314
specific parser function @{ML prop in OuterParse} leads to more instructive
+ − 315
error messages, since the parser will complain that a proposition was expected
+ − 316
when something else than a string or identifier is found.
+ − 317
An optional locale target specification of the form \isa{(\isacommand{in}\ $\ldots$)}
+ − 318
can be parsed using @{ML opt_target in OuterParse}.
+ − 319
The lists of names of the predicates and parameters, together with optional
+ − 320
types and syntax, are parsed using the functions @{ML "fixes" in OuterParse}
+ − 321
and @{ML for_fixes in OuterParse}, respectively.
+ − 322
In addition, the following function from @{ML_struct SpecParse} for parsing
+ − 323
an optional theorem name and attribute, followed by a delimiter, will be useful:
+ − 324
+ − 325
\begin{table}
+ − 326
@{ML "opt_thm_name:
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 327
string -> Attrib.binding parser" in SpecParse}
88
+ − 328
\end{table}
+ − 329
+ − 330
We now have all the necessary tools to write the parser for our
+ − 331
\isa{\isacommand{simple{\isacharunderscore}inductive}} command:
+ − 332
116
+ − 333
88
+ − 334
Once all arguments of the command have been parsed, we apply the function
+ − 335
@{ML add_inductive in SimpleInductivePackage}, which yields a local theory
+ − 336
transformer of type @{ML_type "local_theory -> local_theory"}. Commands in
+ − 337
Isabelle/Isar are realized by transition transformers of type
+ − 338
@{ML_type [display] "Toplevel.transition -> Toplevel.transition"}
+ − 339
We can turn a local theory transformer into a transition transformer by using
+ − 340
the function
+ − 341
+ − 342
@{ML [display] "Toplevel.local_theory : string option ->
32
+ − 343
(local_theory -> local_theory) ->
+ − 344
Toplevel.transition -> Toplevel.transition"}
88
+ − 345
+ − 346
which, apart from the local theory transformer, takes an optional name of a locale
+ − 347
to be used as a basis for the local theory.
60
5b9c6010897b
doem tuning and made the cookbook work again with recent changes (CookBook/Package/Ind_Interface.thy needs to be looked at to see what the problem with the new parser type is)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 348
88
+ − 349
(FIXME : needs to be adjusted to new parser type)
60
5b9c6010897b
doem tuning and made the cookbook work again with recent changes (CookBook/Package/Ind_Interface.thy needs to be looked at to see what the problem with the new parser type is)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 350
88
+ − 351
{\it
+ − 352
The whole parser for our command has type
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
+ − 353
@{text [display] "OuterLex.token list ->
32
+ − 354
(Toplevel.transition -> Toplevel.transition) * OuterLex.token list"}
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
+ − 355
which is abbreviated by @{text OuterSyntax.parser_fn}. The new command can be added
88
+ − 356
to the system via the function
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
+ − 357
@{text [display] "OuterSyntax.command :
32
+ − 358
string -> string -> OuterKeyword.T -> OuterSyntax.parser_fn -> unit"}
88
+ − 359
which imperatively updates the parser table behind the scenes. }
60
5b9c6010897b
doem tuning and made the cookbook work again with recent changes (CookBook/Package/Ind_Interface.thy needs to be looked at to see what the problem with the new parser type is)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 360
88
+ − 361
In addition to the parser, this
+ − 362
function takes two strings representing the name of the command and a short description,
+ − 363
as well as an element of type @{ML_type OuterKeyword.T} describing which \emph{kind} of
+ − 364
command we intend to add. Since we want to add a command for declaring new concepts,
+ − 365
we choose the kind @{ML "OuterKeyword.thy_decl"}. Other kinds include
+ − 366
@{ML "OuterKeyword.thy_goal"}, which is similar to @{ML thy_decl in OuterKeyword},
+ − 367
but requires the user to prove a goal before making the declaration, or
+ − 368
@{ML "OuterKeyword.diag"}, which corresponds to a purely diagnostic command that does
+ − 369
not change the context. For example, the @{ML thy_goal in OuterKeyword} kind is used
+ − 370
by the \isa{\isacommand{function}} command \cite{Krauss-IJCAR06}, which requires the user
+ − 371
to prove that a given set of equations is non-overlapping and covers all cases. The kind
+ − 372
of the command should be chosen with care, since selecting the wrong one can cause strange
+ − 373
behaviour of the user interface, such as failure of the undo mechanism.
32
+ − 374
*}
+ − 375
127
+ − 376
text {*
+ − 377
Note that the @{text trcl} predicate has two different kinds of parameters: the
+ − 378
first parameter @{text R} stays \emph{fixed} throughout the definition, whereas
+ − 379
the second and third parameter changes in the ``recursive call''. This will
+ − 380
become important later on when we deal with fixed parameters and locales.
+ − 381
+ − 382
+ − 383
+ − 384
The purpose of the package we show next is that the user just specifies the
+ − 385
inductive predicate by stating some introduction rules and then the packages
+ − 386
makes the equivalent definition and derives from it the needed properties.
+ − 387
*}
+ − 388
+ − 389
text {*
212
+ − 390
(FIXME: perhaps move somewhere else)
127
+ − 391
212
+ − 392
The point of these examples is to get a feeling what the automatic proofs
+ − 393
should do in order to solve all inductive definitions we throw at them. For this
+ − 394
it is instructive to look at the general construction principle
+ − 395
of inductive definitions, which we shall do in the next section.
215
+ − 396
+ − 397
The code of the inductive package falls roughly in tree parts: the first
+ − 398
deals with the definitions, the second with the induction principles and
+ − 399
the third with the introduction rules.
+ − 400
127
+ − 401
*}
+ − 402
+ − 403
32
+ − 404
(*<*)
+ − 405
end
+ − 406
(*>*)