5 chapter {* Isabelle/HOL \label{ch:hol} *}
7 section {* Primitive types \label{sec:hol-typedef} *}
10 \begin{matharray}{rcl}
11 @{command_def (HOL) "typedecl"} & : & @{text "theory \<rightarrow> theory"} \\
12 @{command_def (HOL) "typedef"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
16 'typedecl' typespec infix?
18 'typedef' altname? abstype '=' repset
21 altname: '(' (name | 'open' | 'open' name) ')'
23 abstype: typespec infix?
25 repset: term ('morphisms' name name)?
31 \item @{command (HOL) "typedecl"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"} is similar
32 to the original @{command "typedecl"} of Isabelle/Pure (see
33 \secref{sec:types-pure}), but also declares type arity @{text "t ::
34 (type, \<dots>, type) type"}, making @{text t} an actual HOL type
35 constructor. %FIXME check, update
37 \item @{command (HOL) "typedef"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = A"} sets up
38 a goal stating non-emptiness of the set @{text A}. After finishing
39 the proof, the theory will be augmented by a Gordon/HOL-style type
40 definition, which establishes a bijection between the representing
41 set @{text A} and the new type @{text t}.
43 Technically, @{command (HOL) "typedef"} defines both a type @{text
44 t} and a set (term constant) of the same name (an alternative base
45 name may be given in parentheses). The injection from type to set
46 is called @{text Rep_t}, its inverse @{text Abs_t} (this may be
47 changed via an explicit @{keyword (HOL) "morphisms"} declaration).
49 Theorems @{text Rep_t}, @{text Rep_t_inverse}, and @{text
50 Abs_t_inverse} provide the most basic characterization as a
51 corresponding injection/surjection pair (in both directions). Rules
52 @{text Rep_t_inject} and @{text Abs_t_inject} provide a slightly
53 more convenient view on the injectivity part, suitable for automated
54 proof tools (e.g.\ in @{attribute simp} or @{attribute iff}
55 declarations). Rules @{text Rep_t_cases}/@{text Rep_t_induct}, and
56 @{text Abs_t_cases}/@{text Abs_t_induct} provide alternative views
57 on surjectivity; these are already declared as set or type rules for
58 the generic @{method cases} and @{method induct} methods.
60 An alternative name may be specified in parentheses; the default is
61 to use @{text t} as indicated before. The ``@{text "(open)"}''
62 declaration suppresses a separate constant definition for the
67 Note that raw type declarations are rarely used in practice; the
68 main application is with experimental (or even axiomatic!) theory
69 fragments. Instead of primitive HOL type definitions, user-level
70 theories usually refer to higher-level packages such as @{command
71 (HOL) "record"} (see \secref{sec:hol-record}) or @{command (HOL)
72 "datatype"} (see \secref{sec:hol-datatype}).
76 section {* Adhoc tuples *}
79 \begin{matharray}{rcl}
80 @{attribute (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\
84 'split\_format' ((( name * ) + 'and') | ('(' 'complete' ')'))
90 \item @{attribute (HOL) split_format}~@{text "p\<^sub>1 \<dots> p\<^sub>m \<AND> \<dots>
91 \<AND> q\<^sub>1 \<dots> q\<^sub>n"} puts expressions of low-level tuple types into
92 canonical form as specified by the arguments given; the @{text i}-th
93 collection of arguments refers to occurrences in premise @{text i}
94 of the rule. The ``@{text "(complete)"}'' option causes \emph{all}
95 arguments in function applications to be represented canonically
96 according to their tuple type structure.
98 Note that these operations tend to invent funny names for new local
99 parameters to be introduced.
105 section {* Records \label{sec:hol-record} *}
108 In principle, records merely generalize the concept of tuples, where
109 components may be addressed by labels instead of just position. The
110 logical infrastructure of records in Isabelle/HOL is slightly more
111 advanced, though, supporting truly extensible record schemes. This
112 admits operations that are polymorphic with respect to record
113 extension, yielding ``object-oriented'' effects like (single)
114 inheritance. See also \cite{NaraschewskiW-TPHOLs98} for more
115 details on object-oriented verification and record subtyping in HOL.
119 subsection {* Basic concepts *}
122 Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records
123 at the level of terms and types. The notation is as follows:
126 \begin{tabular}{l|l|l}
127 & record terms & record types \\ \hline
128 fixed & @{text "\<lparr>x = a, y = b\<rparr>"} & @{text "\<lparr>x :: A, y :: B\<rparr>"} \\
129 schematic & @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} &
130 @{text "\<lparr>x :: A, y :: B, \<dots> :: M\<rparr>"} \\
134 \noindent The ASCII representation of @{text "\<lparr>x = a\<rparr>"} is @{text
137 A fixed record @{text "\<lparr>x = a, y = b\<rparr>"} has field @{text x} of value
138 @{text a} and field @{text y} of value @{text b}. The corresponding
139 type is @{text "\<lparr>x :: A, y :: B\<rparr>"}, assuming that @{text "a :: A"}
140 and @{text "b :: B"}.
142 A record scheme like @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} contains fields
143 @{text x} and @{text y} as before, but also possibly further fields
144 as indicated by the ``@{text "\<dots>"}'' notation (which is actually part
145 of the syntax). The improper field ``@{text "\<dots>"}'' of a record
146 scheme is called the \emph{more part}. Logically it is just a free
147 variable, which is occasionally referred to as ``row variable'' in
148 the literature. The more part of a record scheme may be
149 instantiated by zero or more further components. For example, the
150 previous scheme may get instantiated to @{text "\<lparr>x = a, y = b, z =
151 c, \<dots> = m'\<rparr>"}, where @{text m'} refers to a different more part.
152 Fixed records are special instances of record schemes, where
153 ``@{text "\<dots>"}'' is properly terminated by the @{text "() :: unit"}
154 element. In fact, @{text "\<lparr>x = a, y = b\<rparr>"} is just an abbreviation
155 for @{text "\<lparr>x = a, y = b, \<dots> = ()\<rparr>"}.
157 \medskip Two key observations make extensible records in a simply
158 typed language like HOL work out:
162 \item the more part is internalized, as a free term or type
165 \item field names are externalized, they cannot be accessed within
166 the logic as first-class values.
170 \medskip In Isabelle/HOL record types have to be defined explicitly,
171 fixing their field names and types, and their (optional) parent
172 record. Afterwards, records may be formed using above syntax, while
173 obeying the canonical order of fields as given by their declaration.
174 The record package provides several standard operations like
175 selectors and updates. The common setup for various generic proof
176 tools enable succinct reasoning patterns. See also the Isabelle/HOL
177 tutorial \cite{isabelle-hol-book} for further instructions on using
182 subsection {* Record specifications *}
185 \begin{matharray}{rcl}
186 @{command_def (HOL) "record"} & : & @{text "theory \<rightarrow> theory"} \\
190 'record' typespec '=' (type '+')? (constdecl +)
196 \item @{command (HOL) "record"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t = \<tau> + c\<^sub>1 :: \<sigma>\<^sub>1
197 \<dots> c\<^sub>n :: \<sigma>\<^sub>n"} defines extensible record type @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"},
198 derived from the optional parent record @{text "\<tau>"} by adding new
199 field components @{text "c\<^sub>i :: \<sigma>\<^sub>i"} etc.
201 The type variables of @{text "\<tau>"} and @{text "\<sigma>\<^sub>i"} need to be
202 covered by the (distinct) parameters @{text "\<alpha>\<^sub>1, \<dots>,
203 \<alpha>\<^sub>m"}. Type constructor @{text t} has to be new, while @{text
204 \<tau>} needs to specify an instance of an existing record type. At
205 least one new field @{text "c\<^sub>i"} has to be specified.
206 Basically, field names need to belong to a unique record. This is
207 not a real restriction in practice, since fields are qualified by
208 the record name internally.
210 The parent record specification @{text \<tau>} is optional; if omitted
211 @{text t} becomes a root record. The hierarchy of all records
212 declared within a theory context forms a forest structure, i.e.\ a
213 set of trees starting with a root record each. There is no way to
214 merge multiple parent records!
216 For convenience, @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} is made a
217 type abbreviation for the fixed record type @{text "\<lparr>c\<^sub>1 ::
218 \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n\<rparr>"}, likewise is @{text
219 "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m, \<zeta>) t_scheme"} made an abbreviation for
220 @{text "\<lparr>c\<^sub>1 :: \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n, \<dots> ::
227 subsection {* Record operations *}
230 Any record definition of the form presented above produces certain
231 standard operations. Selectors and updates are provided for any
232 field, including the improper one ``@{text more}''. There are also
233 cumulative record constructor functions. To simplify the
234 presentation below, we assume for now that @{text "(\<alpha>\<^sub>1, \<dots>,
235 \<alpha>\<^sub>m) t"} is a root record with fields @{text "c\<^sub>1 ::
236 \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n"}.
238 \medskip \textbf{Selectors} and \textbf{updates} are available for
239 any field (including ``@{text more}''):
241 \begin{matharray}{lll}
242 @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
243 @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
246 There is special syntax for application of updates: @{text "r\<lparr>x :=
247 a\<rparr>"} abbreviates term @{text "x_update a r"}. Further notation for
248 repeated updates is also available: @{text "r\<lparr>x := a\<rparr>\<lparr>y := b\<rparr>\<lparr>z :=
249 c\<rparr>"} may be written @{text "r\<lparr>x := a, y := b, z := c\<rparr>"}. Note that
250 because of postfix notation the order of fields shown here is
251 reverse than in the actual term. Since repeated updates are just
252 function applications, fields may be freely permuted in @{text "\<lparr>x
253 := a, y := b, z := c\<rparr>"}, as far as logical equality is concerned.
254 Thus commutativity of independent updates can be proven within the
255 logic for any two fields, but not as a general theorem.
257 \medskip The \textbf{make} operation provides a cumulative record
258 constructor function:
260 \begin{matharray}{lll}
261 @{text "t.make"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
264 \medskip We now reconsider the case of non-root records, which are
265 derived of some parent. In general, the latter may depend on
266 another parent as well, resulting in a list of \emph{ancestor
267 records}. Appending the lists of fields of all ancestors results in
268 a certain field prefix. The record package automatically takes care
269 of this by lifting operations over this context of ancestor fields.
270 Assuming that @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} has ancestor
271 fields @{text "b\<^sub>1 :: \<rho>\<^sub>1, \<dots>, b\<^sub>k :: \<rho>\<^sub>k"},
272 the above record operations will get the following types:
276 @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
277 @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow>
278 \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow>
279 \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
280 @{text "t.make"} & @{text "::"} & @{text "\<rho>\<^sub>1 \<Rightarrow> \<dots> \<rho>\<^sub>k \<Rightarrow> \<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow>
281 \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
285 \noindent Some further operations address the extension aspect of a
286 derived record scheme specifically: @{text "t.fields"} produces a
287 record fragment consisting of exactly the new fields introduced here
288 (the result may serve as a more part elsewhere); @{text "t.extend"}
289 takes a fixed record and adds a given more part; @{text
290 "t.truncate"} restricts a record scheme to a fixed record.
294 @{text "t.fields"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
295 @{text "t.extend"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr> \<Rightarrow>
296 \<zeta> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
297 @{text "t.truncate"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
301 \noindent Note that @{text "t.make"} and @{text "t.fields"} coincide
306 subsection {* Derived rules and proof tools *}
309 The record package proves several results internally, declaring
310 these facts to appropriate proof tools. This enables users to
311 reason about record structures quite conveniently. Assume that
312 @{text t} is a record type as specified above.
316 \item Standard conversions for selectors or updates applied to
317 record constructor terms are made part of the default Simplifier
318 context; thus proofs by reduction of basic operations merely require
319 the @{method simp} method without further arguments. These rules
320 are available as @{text "t.simps"}, too.
322 \item Selectors applied to updated records are automatically reduced
323 by an internal simplification procedure, which is also part of the
324 standard Simplifier setup.
326 \item Inject equations of a form analogous to @{prop "(x, y) = (x',
327 y') \<equiv> x = x' \<and> y = y'"} are declared to the Simplifier and Classical
328 Reasoner as @{attribute iff} rules. These rules are available as
331 \item The introduction rule for record equality analogous to @{text
332 "x r = x r' \<Longrightarrow> y r = y r' \<dots> \<Longrightarrow> r = r'"} is declared to the Simplifier,
333 and as the basic rule context as ``@{attribute intro}@{text "?"}''.
334 The rule is called @{text "t.equality"}.
336 \item Representations of arbitrary record expressions as canonical
337 constructor terms are provided both in @{method cases} and @{method
338 induct} format (cf.\ the generic proof methods of the same name,
339 \secref{sec:cases-induct}). Several variations are available, for
340 fixed records, record schemes, more parts etc.
342 The generic proof methods are sufficiently smart to pick the most
343 sensible rule according to the type of the indicated record
344 expression: users just need to apply something like ``@{text "(cases
345 r)"}'' to a certain proof problem.
347 \item The derived record operations @{text "t.make"}, @{text
348 "t.fields"}, @{text "t.extend"}, @{text "t.truncate"} are \emph{not}
349 treated automatically, but usually need to be expanded by hand,
350 using the collective fact @{text "t.defs"}.
356 section {* Datatypes \label{sec:hol-datatype} *}
359 \begin{matharray}{rcl}
360 @{command_def (HOL) "datatype"} & : & @{text "theory \<rightarrow> theory"} \\
361 @{command_def (HOL) "rep_datatype"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
365 'datatype' (dtspec + 'and')
367 'rep\_datatype' ('(' (name +) ')')? (term +)
370 dtspec: parname? typespec infix? '=' (cons + '|')
372 cons: name ( type * ) mixfix?
377 \item @{command (HOL) "datatype"} defines inductive datatypes in
380 \item @{command (HOL) "rep_datatype"} represents existing types as
381 inductive ones, generating the standard infrastructure of derived
382 concepts (primitive recursion etc.).
386 The induction and exhaustion theorems generated provide case names
387 according to the constructors involved, while parameters are named
388 after the types (see also \secref{sec:cases-induct}).
390 See \cite{isabelle-HOL} for more details on datatypes, but beware of
391 the old-style theory syntax being used there! Apart from proper
392 proof methods for case-analysis and induction, there are also
393 emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL)
394 induct_tac} available, see \secref{sec:hol-induct-tac}; these admit
395 to refer directly to the internal structure of subgoals (including
396 internally bound parameters).
400 section {* Recursive functions \label{sec:recursion} *}
403 \begin{matharray}{rcl}
404 @{command_def (HOL) "primrec"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
405 @{command_def (HOL) "fun"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
406 @{command_def (HOL) "function"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
407 @{command_def (HOL) "termination"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
411 'primrec' target? fixes 'where' equations
413 equations: (thmdecl? prop + '|')
415 ('fun' | 'function') target? functionopts? fixes 'where' clauses
417 clauses: (thmdecl? prop ('(' 'otherwise' ')')? + '|')
419 functionopts: '(' (('sequential' | 'domintros' | 'tailrec' | 'default' term) + ',') ')'
421 'termination' ( term )?
426 \item @{command (HOL) "primrec"} defines primitive recursive
427 functions over datatypes, see also \cite{isabelle-HOL}.
429 \item @{command (HOL) "function"} defines functions by general
430 wellfounded recursion. A detailed description with examples can be
431 found in \cite{isabelle-function}. The function is specified by a
432 set of (possibly conditional) recursive equations with arbitrary
433 pattern matching. The command generates proof obligations for the
434 completeness and the compatibility of patterns.
436 The defined function is considered partial, and the resulting
437 simplification rules (named @{text "f.psimps"}) and induction rule
438 (named @{text "f.pinduct"}) are guarded by a generated domain
439 predicate @{text "f_dom"}. The @{command (HOL) "termination"}
440 command can then be used to establish that the function is total.
442 \item @{command (HOL) "fun"} is a shorthand notation for ``@{command
443 (HOL) "function"}~@{text "(sequential)"}, followed by automated
444 proof attempts regarding pattern matching and termination. See
445 \cite{isabelle-function} for further details.
447 \item @{command (HOL) "termination"}~@{text f} commences a
448 termination proof for the previously defined function @{text f}. If
449 this is omitted, the command refers to the most recent function
450 definition. After the proof is closed, the recursive equations and
451 the induction principle is established.
455 Recursive definitions introduced by the @{command (HOL) "function"}
457 reasoning by induction (cf.\ \secref{sec:cases-induct}): rule @{text
458 "c.induct"} (where @{text c} is the name of the function definition)
459 refers to a specific induction rule, with parameters named according
460 to the user-specified equations. Cases are numbered (starting from 1).
462 For @{command (HOL) "primrec"}, the induction principle coincides
463 with structural recursion on the datatype the recursion is carried
466 The equations provided by these packages may be referred later as
467 theorem list @{text "f.simps"}, where @{text f} is the (collective)
468 name of the functions defined. Individual equations may be named
471 The @{command (HOL) "function"} command accepts the following
476 \item @{text sequential} enables a preprocessor which disambiguates
477 overlapping patterns by making them mutually disjoint. Earlier
478 equations take precedence over later ones. This allows to give the
479 specification in a format very similar to functional programming.
480 Note that the resulting simplification and induction rules
481 correspond to the transformed specification, not the one given
482 originally. This usually means that each equation given by the user
483 may result in several theroems. Also note that this automatic
484 transformation only works for ML-style datatype patterns.
486 \item @{text domintros} enables the automated generation of
487 introduction rules for the domain predicate. While mostly not
488 needed, they can be helpful in some proofs about partial functions.
490 \item @{text tailrec} generates the unconstrained recursive
491 equations even without a termination proof, provided that the
492 function is tail-recursive. This currently only works
494 \item @{text "default d"} allows to specify a default value for a
495 (partial) function, which will ensure that @{text "f x = d x"}
496 whenever @{text "x \<notin> f_dom"}.
502 subsection {* Proof methods related to recursive definitions *}
505 \begin{matharray}{rcl}
506 @{method_def (HOL) pat_completeness} & : & @{text method} \\
507 @{method_def (HOL) relation} & : & @{text method} \\
508 @{method_def (HOL) lexicographic_order} & : & @{text method} \\
509 @{method_def (HOL) size_change} & : & @{text method} \\
515 'lexicographic\_order' ( clasimpmod * )
517 'size\_change' ( orders ( clasimpmod * ) )
519 orders: ( 'max' | 'min' | 'ms' ) *
524 \item @{method (HOL) pat_completeness} is a specialized method to
525 solve goals regarding the completeness of pattern matching, as
526 required by the @{command (HOL) "function"} package (cf.\
527 \cite{isabelle-function}).
529 \item @{method (HOL) relation}~@{text R} introduces a termination
530 proof using the relation @{text R}. The resulting proof state will
531 contain goals expressing that @{text R} is wellfounded, and that the
532 arguments of recursive calls decrease with respect to @{text R}.
533 Usually, this method is used as the initial proof step of manual
536 \item @{method (HOL) "lexicographic_order"} attempts a fully
537 automated termination proof by searching for a lexicographic
538 combination of size measures on the arguments of the function. The
539 method accepts the same arguments as the @{method auto} method,
540 which it uses internally to prove local descents. The same context
541 modifiers as for @{method auto} are accepted, see
542 \secref{sec:clasimp}.
544 In case of failure, extensive information is printed, which can help
545 to analyse the situation (cf.\ \cite{isabelle-function}).
547 \item @{method (HOL) "size_change"} also works on termination goals,
548 using a variation of the size-change principle, together with a
549 graph decomposition technique (see \cite{krauss_phd} for details).
550 Three kinds of orders are used internally: @{text max}, @{text min},
551 and @{text ms} (multiset), which is only available when the theory
552 @{text Multiset} is loaded. When no order kinds are given, they are
553 tried in order. The search for a termination proof uses SAT solving
556 For local descent proofs, the same context modifiers as for @{method
557 auto} are accepted, see \secref{sec:clasimp}.
563 subsection {* Old-style recursive function definitions (TFL) *}
566 The old TFL commands @{command (HOL) "recdef"} and @{command (HOL)
567 "recdef_tc"} for defining recursive are mostly obsolete; @{command
568 (HOL) "function"} or @{command (HOL) "fun"} should be used instead.
570 \begin{matharray}{rcl}
571 @{command_def (HOL) "recdef"} & : & @{text "theory \<rightarrow> theory)"} \\
572 @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
576 'recdef' ('(' 'permissive' ')')? \\ name term (prop +) hints?
580 hints: '(' 'hints' ( recdefmod * ) ')'
582 recdefmod: (('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del') ':' thmrefs) | clasimpmod
584 tc: nameref ('(' nat ')')?
590 \item @{command (HOL) "recdef"} defines general well-founded
591 recursive functions (using the TFL package), see also
592 \cite{isabelle-HOL}. The ``@{text "(permissive)"}'' option tells
593 TFL to recover from failed proof attempts, returning unfinished
594 results. The @{text recdef_simp}, @{text recdef_cong}, and @{text
595 recdef_wf} hints refer to auxiliary rules to be used in the internal
596 automated proof process of TFL. Additional @{syntax clasimpmod}
597 declarations (cf.\ \secref{sec:clasimp}) may be given to tune the
598 context of the Simplifier (cf.\ \secref{sec:simplifier}) and
599 Classical reasoner (cf.\ \secref{sec:classical}).
601 \item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the
602 proof for leftover termination condition number @{text i} (default
603 1) as generated by a @{command (HOL) "recdef"} definition of
606 Note that in most cases, @{command (HOL) "recdef"} is able to finish
607 its internal proofs without manual intervention.
611 \medskip Hints for @{command (HOL) "recdef"} may be also declared
612 globally, using the following attributes.
614 \begin{matharray}{rcl}
615 @{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\
616 @{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\
617 @{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\
621 ('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del')
627 section {* Inductive and coinductive definitions \label{sec:hol-inductive} *}
630 An \textbf{inductive definition} specifies the least predicate (or
631 set) @{text R} closed under given rules: applying a rule to elements
632 of @{text R} yields a result within @{text R}. For example, a
633 structural operational semantics is an inductive definition of an
636 Dually, a \textbf{coinductive definition} specifies the greatest
637 predicate~/ set @{text R} that is consistent with given rules: every
638 element of @{text R} can be seen as arising by applying a rule to
639 elements of @{text R}. An important example is using bisimulation
640 relations to formalise equivalence of processes and infinite data
643 \medskip The HOL package is related to the ZF one, which is
644 described in a separate paper,\footnote{It appeared in CADE
645 \cite{paulson-CADE}; a longer version is distributed with Isabelle.}
646 which you should refer to in case of difficulties. The package is
647 simpler than that of ZF thanks to implicit type-checking in HOL.
648 The types of the (co)inductive predicates (or sets) determine the
649 domain of the fixedpoint definition, and the package does not have
650 to use inference rules for type-checking.
652 \begin{matharray}{rcl}
653 @{command_def (HOL) "inductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
654 @{command_def (HOL) "inductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
655 @{command_def (HOL) "coinductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
656 @{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
657 @{attribute_def (HOL) mono} & : & @{text attribute} \\
661 ('inductive' | 'inductive\_set' | 'coinductive' | 'coinductive\_set') target? fixes ('for' fixes)? \\
662 ('where' clauses)? ('monos' thmrefs)?
664 clauses: (thmdecl? prop + '|')
666 'mono' (() | 'add' | 'del')
672 \item @{command (HOL) "inductive"} and @{command (HOL)
673 "coinductive"} define (co)inductive predicates from the
674 introduction rules given in the @{keyword "where"} part. The
675 optional @{keyword "for"} part contains a list of parameters of the
676 (co)inductive predicates that remain fixed throughout the
677 definition. The optional @{keyword "monos"} section contains
678 \emph{monotonicity theorems}, which are required for each operator
679 applied to a recursive set in the introduction rules. There
680 \emph{must} be a theorem of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"},
681 for each premise @{text "M R\<^sub>i t"} in an introduction rule!
683 \item @{command (HOL) "inductive_set"} and @{command (HOL)
684 "coinductive_set"} are wrappers for to the previous commands,
685 allowing the definition of (co)inductive sets.
687 \item @{attribute (HOL) mono} declares monotonicity rules. These
688 rule are involved in the automated monotonicity proof of @{command
695 subsection {* Derived rules *}
698 Each (co)inductive definition @{text R} adds definitions to the
699 theory and also proves some theorems:
703 \item @{text R.intros} is the list of introduction rules as proven
704 theorems, for the recursive predicates (or sets). The rules are
705 also available individually, using the names given them in the
708 \item @{text R.cases} is the case analysis (or elimination) rule;
710 \item @{text R.induct} or @{text R.coinduct} is the (co)induction
715 When several predicates @{text "R\<^sub>1, \<dots>, R\<^sub>n"} are
716 defined simultaneously, the list of introduction rules is called
717 @{text "R\<^sub>1_\<dots>_R\<^sub>n.intros"}, the case analysis rules are
718 called @{text "R\<^sub>1.cases, \<dots>, R\<^sub>n.cases"}, and the list
719 of mutual induction rules is called @{text
720 "R\<^sub>1_\<dots>_R\<^sub>n.inducts"}.
724 subsection {* Monotonicity theorems *}
727 Each theory contains a default set of theorems that are used in
728 monotonicity proofs. New rules can be added to this set via the
729 @{attribute (HOL) mono} attribute. The HOL theory @{text Inductive}
730 shows how this is done. In general, the following monotonicity
731 theorems may be added:
735 \item Theorems of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"}, for proving
736 monotonicity of inductive definitions whose introduction rules have
737 premises involving terms such as @{text "M R\<^sub>i t"}.
739 \item Monotonicity theorems for logical operators, which are of the
740 general form @{text "(\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> (\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> \<longrightarrow> \<dots>"}. For example, in
741 the case of the operator @{text "\<or>"}, the corresponding theorem is
743 \infer{@{text "P\<^sub>1 \<or> P\<^sub>2 \<longrightarrow> Q\<^sub>1 \<or> Q\<^sub>2"}}{@{text "P\<^sub>1 \<longrightarrow> Q\<^sub>1"} & @{text "P\<^sub>2 \<longrightarrow> Q\<^sub>2"}}
746 \item De Morgan style equations for reasoning about the ``polarity''
749 @{prop "\<not> \<not> P \<longleftrightarrow> P"} \qquad\qquad
750 @{prop "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"}
753 \item Equations for reducing complex operators to more primitive
754 ones whose monotonicity can easily be proved, e.g.
756 @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> \<not> P \<or> Q"} \qquad\qquad
757 @{prop "Ball A P \<equiv> \<forall>x. x \<in> A \<longrightarrow> P x"}
762 %FIXME: Example of an inductive definition
766 section {* Arithmetic proof support *}
769 \begin{matharray}{rcl}
770 @{method_def (HOL) arith} & : & @{text method} \\
771 @{attribute_def (HOL) arith} & : & @{text attribute} \\
772 @{attribute_def (HOL) arith_split} & : & @{text attribute} \\
775 The @{method (HOL) arith} method decides linear arithmetic problems
776 (on types @{text nat}, @{text int}, @{text real}). Any current
777 facts are inserted into the goal before running the procedure.
779 The @{attribute (HOL) arith} attribute declares facts that are
780 always supplied to the arithmetic provers implicitly.
782 The @{attribute (HOL) arith_split} attribute declares case split
783 rules to be expanded before @{method (HOL) arith} is invoked.
785 Note that a simpler (but faster) arithmetic prover is
786 already invoked by the Simplifier.
790 section {* Intuitionistic proof search *}
793 \begin{matharray}{rcl}
794 @{method_def (HOL) iprover} & : & @{text method} \\
798 'iprover' ('!' ?) ( rulemod * )
802 The @{method (HOL) iprover} method performs intuitionistic proof
803 search, depending on specifically declared rules from the context,
804 or given as explicit arguments. Chained facts are inserted into the
805 goal before commencing proof search; ``@{method (HOL) iprover}@{text
806 "!"}'' means to include the current @{fact prems} as well.
808 Rules need to be classified as @{attribute (Pure) intro},
809 @{attribute (Pure) elim}, or @{attribute (Pure) dest}; here the
810 ``@{text "!"}'' indicator refers to ``safe'' rules, which may be
811 applied aggressively (without considering back-tracking later).
812 Rules declared with ``@{text "?"}'' are ignored in proof search (the
813 single-step @{method rule} method still observes these). An
814 explicit weight annotation may be given as well; otherwise the
815 number of rule premises will be taken into account here.
819 section {* Coherent Logic *}
822 \begin{matharray}{rcl}
823 @{method_def (HOL) "coherent"} & : & @{text method} \\
831 The @{method (HOL) coherent} method solves problems of
832 \emph{Coherent Logic} \cite{Bezem-Coquand:2005}, which covers
833 applications in confluence theory, lattice theory and projective
834 geometry. See @{"file" "~~/src/HOL/ex/Coherent.thy"} for some
839 section {* Checking and refuting propositions *}
842 Identifying incorrect propositions usually involves evaluation of
843 particular assignments and systematic counter example search. This
844 is supported by the following commands.
846 \begin{matharray}{rcl}
847 @{command_def (HOL) "value"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
848 @{command_def (HOL) "quickcheck"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
849 @{command_def (HOL) "quickcheck_params"} & : & @{text "theory \<rightarrow> theory"}
853 'value' ( ( '[' name ']' ) ? ) modes? term
856 'quickcheck' ( ( '[' args ']' ) ? ) nat?
859 'quickcheck_params' ( ( '[' args ']' ) ? )
862 modes: '(' (name + ) ')'
865 args: ( name '=' value + ',' )
871 \item @{command (HOL) "value"}~@{text t} evaluates and prints a
872 term; optionally @{text modes} can be specified, which are
873 appended to the current print mode (see also \cite{isabelle-ref}).
874 Internally, the evaluation is performed by registered evaluators,
875 which are invoked sequentially until a result is returned.
876 Alternatively a specific evaluator can be selected using square
877 brackets; available evaluators include @{text nbe} for
878 \emph{normalization by evaluation} and \emph{code} for code
881 \item @{command (HOL) "quickcheck"} tests the current goal for
882 counter examples using a series of arbitrary assignments for its
883 free variables; by default the first subgoal is tested, an other
884 can be selected explicitly using an optional goal index.
885 A number of configuration options are supported for
886 @{command (HOL) "quickcheck"}, notably:
890 \item[size] specifies the maximum size of the search space for
893 \item[iterations] sets how many sets of assignments are
894 generated for each particular size.
896 \item[no\_assms] specifies whether assumptions in
897 structured proofs should be ignored.
901 These option can be given within square brackets.
903 \item @{command (HOL) "quickcheck_params"} changes quickcheck
904 configuration options persitently.
910 section {* Invoking automated reasoning tools -- The Sledgehammer *}
913 Isabelle/HOL includes a generic \emph{ATP manager} that allows
914 external automated reasoning tools to crunch a pending goal.
915 Supported provers include E\footnote{\url{http://www.eprover.org}},
916 SPASS\footnote{\url{http://www.spass-prover.org/}}, and Vampire.
917 There is also a wrapper to invoke provers remotely via the
918 SystemOnTPTP\footnote{\url{http://www.cs.miami.edu/~tptp/cgi-bin/SystemOnTPTP}}
921 The problem passed to external provers consists of the goal together
922 with a smart selection of lemmas from the current theory context.
923 The result of a successful proof search is some source text that
924 usually reconstructs the proof within Isabelle, without requiring
925 external provers again. The Metis
926 prover\footnote{\url{http://www.gilith.com/software/metis/}} that is
927 integrated into Isabelle/HOL is being used here.
929 In this mode of operation, heavy means of automated reasoning are
930 used as a strong relevance filter, while the main proof checking
931 works via explicit inferences going through the Isabelle kernel.
932 Moreover, rechecking Isabelle proof texts with already specified
933 auxiliary facts is much faster than performing fully automated
934 search over and over again.
936 \begin{matharray}{rcl}
937 @{command_def (HOL) "sledgehammer"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
938 @{command_def (HOL) "print_atps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
939 @{command_def (HOL) "atp_info"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
940 @{command_def (HOL) "atp_kill"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
941 @{command_def (HOL) "atp_messages"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
942 @{method_def (HOL) metis} & : & @{text method} \\
946 'sledgehammer' ( nameref * )
948 'atp\_messages' ('(' nat ')')?
957 \item @{command (HOL) sledgehammer}~@{text "prover\<^sub>1 \<dots> prover\<^sub>n"}
958 invokes the specified automated theorem provers on the first
959 subgoal. Provers are run in parallel, the first successful result
960 is displayed, and the other attempts are terminated.
962 Provers are defined in the theory context, see also @{command (HOL)
963 print_atps}. If no provers are given as arguments to @{command
964 (HOL) sledgehammer}, the system refers to the default defined as
965 ``ATP provers'' preference by the user interface.
967 There are additional preferences for timeout (default: 60 seconds),
968 and the maximum number of independent prover processes (default: 5);
969 excessive provers are automatically terminated.
971 \item @{command (HOL) print_atps} prints the list of automated
972 theorem provers available to the @{command (HOL) sledgehammer}
975 \item @{command (HOL) atp_info} prints information about presently
976 running provers, including elapsed runtime, and the remaining time
979 \item @{command (HOL) atp_kill} terminates all presently running
982 \item @{command (HOL) atp_messages} displays recent messages issued
983 by automated theorem provers. This allows to examine results that
984 might have got lost due to the asynchronous nature of default
985 @{command (HOL) sledgehammer} output. An optional message limit may
986 be specified (default 5).
988 \item @{method (HOL) metis}~@{text "facts"} invokes the Metis prover
989 with the given facts. Metis is an automated proof tool of medium
990 strength, but is fully integrated into Isabelle/HOL, with explicit
991 inferences going through the kernel. Thus its results are
992 guaranteed to be ``correct by construction''.
994 Note that all facts used with Metis need to be specified as explicit
995 arguments. There are no rule declarations as for other Isabelle
996 provers, like @{method blast} or @{method fast}.
1002 section {* Unstructured case analysis and induction \label{sec:hol-induct-tac} *}
1005 The following tools of Isabelle/HOL support cases analysis and
1006 induction in unstructured tactic scripts; see also
1007 \secref{sec:cases-induct} for proper Isar versions of similar ideas.
1009 \begin{matharray}{rcl}
1010 @{method_def (HOL) case_tac}@{text "\<^sup>*"} & : & @{text method} \\
1011 @{method_def (HOL) induct_tac}@{text "\<^sup>*"} & : & @{text method} \\
1012 @{method_def (HOL) ind_cases}@{text "\<^sup>*"} & : & @{text method} \\
1013 @{command_def (HOL) "inductive_cases"}@{text "\<^sup>*"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
1017 'case\_tac' goalspec? term rule?
1019 'induct\_tac' goalspec? (insts * 'and') rule?
1021 'ind\_cases' (prop +) ('for' (name +)) ?
1023 'inductive\_cases' (thmdecl? (prop +) + 'and')
1026 rule: ('rule' ':' thmref)
1032 \item @{method (HOL) case_tac} and @{method (HOL) induct_tac} admit
1033 to reason about inductive types. Rules are selected according to
1034 the declarations by the @{attribute cases} and @{attribute induct}
1035 attributes, cf.\ \secref{sec:cases-induct}. The @{command (HOL)
1036 datatype} package already takes care of this.
1038 These unstructured tactics feature both goal addressing and dynamic
1039 instantiation. Note that named rule cases are \emph{not} provided
1040 as would be by the proper @{method cases} and @{method induct} proof
1041 methods (see \secref{sec:cases-induct}). Unlike the @{method
1042 induct} method, @{method induct_tac} does not handle structured rule
1043 statements, only the compact object-logic conclusion of the subgoal
1046 \item @{method (HOL) ind_cases} and @{command (HOL)
1047 "inductive_cases"} provide an interface to the internal @{ML_text
1048 mk_cases} operation. Rules are simplified in an unrestricted
1051 While @{method (HOL) ind_cases} is a proof method to apply the
1052 result immediately as elimination rules, @{command (HOL)
1053 "inductive_cases"} provides case split theorems at the theory level
1054 for later use. The @{keyword "for"} argument of the @{method (HOL)
1055 ind_cases} method allows to specify a list of variables that should
1056 be generalized before applying the resulting rule.
1062 section {* Executable code *}
1065 Isabelle/Pure provides two generic frameworks to support code
1066 generation from executable specifications. Isabelle/HOL
1067 instantiates these mechanisms in a way that is amenable to end-user
1070 One framework generates code from both functional and relational
1071 programs to SML. See \cite{isabelle-HOL} for further information
1072 (this actually covers the new-style theory format as well).
1074 \begin{matharray}{rcl}
1075 @{command_def (HOL) "code_module"} & : & @{text "theory \<rightarrow> theory"} \\
1076 @{command_def (HOL) "code_library"} & : & @{text "theory \<rightarrow> theory"} \\
1077 @{command_def (HOL) "consts_code"} & : & @{text "theory \<rightarrow> theory"} \\
1078 @{command_def (HOL) "types_code"} & : & @{text "theory \<rightarrow> theory"} \\
1079 @{attribute_def (HOL) code} & : & @{text attribute} \\
1083 ( 'code\_module' | 'code\_library' ) modespec ? name ? \\
1084 ( 'file' name ) ? ( 'imports' ( name + ) ) ? \\
1085 'contains' ( ( name '=' term ) + | term + )
1088 modespec: '(' ( name * ) ')'
1091 'consts\_code' (codespec +)
1094 codespec: const template attachment ?
1097 'types\_code' (tycodespec +)
1100 tycodespec: name template attachment ?
1106 template: '(' string ')'
1109 attachment: 'attach' modespec ? verblbrace text verbrbrace
1116 \medskip The other framework generates code from functional programs
1117 (including overloading using type classes) to SML \cite{SML}, OCaml
1118 \cite{OCaml} and Haskell \cite{haskell-revised-report}.
1119 Conceptually, code generation is split up in three steps:
1120 \emph{selection} of code theorems, \emph{translation} into an
1121 abstract executable view and \emph{serialization} to a specific
1122 \emph{target language}. See \cite{isabelle-codegen} for an
1123 introduction on how to use it.
1125 \begin{matharray}{rcl}
1126 @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
1127 @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
1128 @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
1129 @{command_def (HOL) "code_datatype"} & : & @{text "theory \<rightarrow> theory"} \\
1130 @{command_def (HOL) "code_const"} & : & @{text "theory \<rightarrow> theory"} \\
1131 @{command_def (HOL) "code_type"} & : & @{text "theory \<rightarrow> theory"} \\
1132 @{command_def (HOL) "code_class"} & : & @{text "theory \<rightarrow> theory"} \\
1133 @{command_def (HOL) "code_instance"} & : & @{text "theory \<rightarrow> theory"} \\
1134 @{command_def (HOL) "code_monad"} & : & @{text "theory \<rightarrow> theory"} \\
1135 @{command_def (HOL) "code_reserved"} & : & @{text "theory \<rightarrow> theory"} \\
1136 @{command_def (HOL) "code_include"} & : & @{text "theory \<rightarrow> theory"} \\
1137 @{command_def (HOL) "code_modulename"} & : & @{text "theory \<rightarrow> theory"} \\
1138 @{command_def (HOL) "code_abort"} & : & @{text "theory \<rightarrow> theory"} \\
1139 @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
1140 @{command_def (HOL) "print_codeproc"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
1141 @{attribute_def (HOL) code} & : & @{text attribute} \\
1145 'export\_code' ( constexpr + ) \\
1146 ( ( 'in' target ( 'module\_name' string ) ? \\
1147 ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
1150 'code\_thms' ( constexpr + ) ?
1153 'code\_deps' ( constexpr + ) ?
1159 constexpr: ( const | 'name.*' | '*' )
1162 typeconstructor: nameref
1168 target: 'OCaml' | 'SML' | 'Haskell'
1171 'code\_datatype' const +
1174 'code\_const' (const + 'and') \\
1175 ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
1178 'code\_type' (typeconstructor + 'and') \\
1179 ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
1182 'code\_class' (class + 'and') \\
1183 ( ( '(' target \\ ( string ? + 'and' ) ')' ) + )
1186 'code\_instance' (( typeconstructor '::' class ) + 'and') \\
1187 ( ( '(' target ( '-' ? + 'and' ) ')' ) + )
1190 'code\_monad' const const target
1193 'code\_reserved' target ( string + )
1196 'code\_include' target ( string ( string | '-') )
1199 'code\_modulename' target ( ( string string ) + )
1202 'code\_abort' ( const + )
1205 syntax: string | ( 'infix' | 'infixl' | 'infixr' ) nat string
1211 'code_unfold' ( 'del' ) ?
1214 'code_post' ( 'del' ) ?
1220 \item @{command (HOL) "export_code"} is the canonical interface for
1221 generating and serializing code: for a given list of constants, code
1222 is generated for the specified target languages. If no serialization
1223 instruction is given, only abstract code is generated internally.
1225 Constants may be specified by giving them literally, referring to
1226 all executable contants within a certain theory by giving @{text
1227 "name.*"}, or referring to \emph{all} executable constants currently
1228 available by giving @{text "*"}.
1230 By default, for each involved theory one corresponding name space
1231 module is generated. Alternativly, a module name may be specified
1232 after the @{keyword "module_name"} keyword; then \emph{all} code is
1233 placed in this module.
1235 For \emph{SML} and \emph{OCaml}, the file specification refers to a
1236 single file; for \emph{Haskell}, it refers to a whole directory,
1237 where code is generated in multiple files reflecting the module
1238 hierarchy. The file specification ``@{text "-"}'' denotes standard
1239 output. For \emph{SML}, omitting the file specification compiles
1240 code internally in the context of the current ML session.
1242 Serializers take an optional list of arguments in parentheses. For
1243 \emph{SML} and \emph{OCaml}, ``@{text no_signatures}`` omits
1244 explicit module signatures.
1246 For \emph{Haskell} a module name prefix may be given using the ``@{text
1247 "root:"}'' argument; ``@{text string_classes}'' adds a ``@{verbatim
1248 "deriving (Read, Show)"}'' clause to each appropriate datatype
1251 \item @{command (HOL) "code_thms"} prints a list of theorems
1252 representing the corresponding program containing all given
1255 \item @{command (HOL) "code_deps"} visualizes dependencies of
1256 theorems representing the corresponding program containing all given
1259 \item @{command (HOL) "code_datatype"} specifies a constructor set
1262 \item @{command (HOL) "code_const"} associates a list of constants
1263 with target-specific serializations; omitting a serialization
1264 deletes an existing serialization.
1266 \item @{command (HOL) "code_type"} associates a list of type
1267 constructors with target-specific serializations; omitting a
1268 serialization deletes an existing serialization.
1270 \item @{command (HOL) "code_class"} associates a list of classes
1271 with target-specific class names; omitting a serialization deletes
1272 an existing serialization. This applies only to \emph{Haskell}.
1274 \item @{command (HOL) "code_instance"} declares a list of type
1275 constructor / class instance relations as ``already present'' for a
1276 given target. Omitting a ``@{text "-"}'' deletes an existing
1277 ``already present'' declaration. This applies only to
1280 \item @{command (HOL) "code_monad"} provides an auxiliary mechanism
1281 to generate monadic code for Haskell.
1283 \item @{command (HOL) "code_reserved"} declares a list of names as
1284 reserved for a given target, preventing it to be shadowed by any
1287 \item @{command (HOL) "code_include"} adds arbitrary named content
1288 (``include'') to generated code. A ``@{text "-"}'' as last argument
1289 will remove an already added ``include''.
1291 \item @{command (HOL) "code_modulename"} declares aliasings from one
1292 module name onto another.
1294 \item @{command (HOL) "code_abort"} declares constants which are not
1295 required to have a definition by means of code equations; if
1296 needed these are implemented by program abort instead.
1298 \item @{attribute (HOL) code} explicitly selects (or with option
1299 ``@{text "del"}'' deselects) a code equation for code
1300 generation. Usually packages introducing code equations provide
1301 a reasonable default setup for selection.
1303 \item @{attribute (HOL) code_inline} declares (or with
1304 option ``@{text "del"}'' removes) inlining theorems which are
1305 applied as rewrite rules to any code equation during
1308 \item @{attribute (HOL) code_post} declares (or with
1309 option ``@{text "del"}'' removes) theorems which are
1310 applied as rewrite rules to any result of an evaluation.
1312 \item @{command (HOL) "print_codesetup"} gives an overview on
1313 selected code equations and code generator datatypes.
1315 \item @{command (HOL) "print_codeproc"} prints the setup
1316 of the code generator preprocessor.
1322 section {* Definition by specification \label{sec:hol-specification} *}
1325 \begin{matharray}{rcl}
1326 @{command_def (HOL) "specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
1327 @{command_def (HOL) "ax_specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
1331 ('specification' | 'ax\_specification') '(' (decl +) ')' \\ (thmdecl? prop +)
1333 decl: ((name ':')? term '(' 'overloaded' ')'?)
1338 \item @{command (HOL) "specification"}~@{text "decls \<phi>"} sets up a
1339 goal stating the existence of terms with the properties specified to
1340 hold for the constants given in @{text decls}. After finishing the
1341 proof, the theory will be augmented with definitions for the given
1342 constants, as well as with theorems stating the properties for these
1345 \item @{command (HOL) "ax_specification"}~@{text "decls \<phi>"} sets up
1346 a goal stating the existence of terms with the properties specified
1347 to hold for the constants given in @{text decls}. After finishing
1348 the proof, the theory will be augmented with axioms expressing the
1349 properties given in the first place.
1351 \item @{text decl} declares a constant to be defined by the
1352 specification given. The definition for the constant @{text c} is
1353 bound to the name @{text c_def} unless a theorem name is given in
1354 the declaration. Overloaded constants should be declared as such.
1358 Whether to use @{command (HOL) "specification"} or @{command (HOL)
1359 "ax_specification"} is to some extent a matter of style. @{command
1360 (HOL) "specification"} introduces no new axioms, and so by
1361 construction cannot introduce inconsistencies, whereas @{command
1362 (HOL) "ax_specification"} does introduce axioms, but only after the
1363 user has explicitly proven it to be safe. A practical issue must be
1364 considered, though: After introducing two constants with the same
1365 properties using @{command (HOL) "specification"}, one can prove
1366 that the two constants are, in fact, equal. If this might be a
1367 problem, one should use @{command (HOL) "ax_specification"}.