7 chapter {* Isabelle/HOL \label{ch:hol} *}
9 section {* Primitive types \label{sec:hol-typedef} *}
12 \begin{matharray}{rcl}
13 @{command_def (HOL) "typedecl"} & : & \isartrans{theory}{theory} \\
14 @{command_def (HOL) "typedef"} & : & \isartrans{theory}{proof(prove)} \\
18 'typedecl' typespec infix?
20 'typedef' altname? abstype '=' repset
23 altname: '(' (name | 'open' | 'open' name) ')'
25 abstype: typespec infix?
27 repset: term ('morphisms' name name)?
33 \item [@{command (HOL) "typedecl"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n)
34 t"}] is similar to the original @{command "typedecl"} of
35 Isabelle/Pure (see \secref{sec:types-pure}), but also declares type
36 arity @{text "t :: (type, \<dots>, type) type"}, making @{text t} an
37 actual HOL type constructor. %FIXME check, update
39 \item [@{command (HOL) "typedef"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n)
40 t = A"}] sets up a goal stating non-emptiness of the set @{text A}.
41 After finishing the proof, the theory will be augmented by a
42 Gordon/HOL-style type definition, which establishes a bijection
43 between the representing set @{text A} and the new type @{text t}.
45 Technically, @{command (HOL) "typedef"} defines both a type @{text
46 t} and a set (term constant) of the same name (an alternative base
47 name may be given in parentheses). The injection from type to set
48 is called @{text Rep_t}, its inverse @{text Abs_t} (this may be
49 changed via an explicit @{keyword (HOL) "morphisms"} declaration).
51 Theorems @{text Rep_t}, @{text Rep_t_inverse}, and @{text
52 Abs_t_inverse} provide the most basic characterization as a
53 corresponding injection/surjection pair (in both directions). Rules
54 @{text Rep_t_inject} and @{text Abs_t_inject} provide a slightly
55 more convenient view on the injectivity part, suitable for automated
56 proof tools (e.g.\ in @{attribute simp} or @{attribute iff}
57 declarations). Rules @{text Rep_t_cases}/@{text Rep_t_induct}, and
58 @{text Abs_t_cases}/@{text Abs_t_induct} provide alternative views
59 on surjectivity; these are already declared as set or type rules for
60 the generic @{method cases} and @{method induct} methods.
62 An alternative name may be specified in parentheses; the default is
63 to use @{text t} as indicated before. The ``@{text "(open)"}''
64 declaration suppresses a separate constant definition for the
69 Note that raw type declarations are rarely used in practice; the
70 main application is with experimental (or even axiomatic!) theory
71 fragments. Instead of primitive HOL type definitions, user-level
72 theories usually refer to higher-level packages such as @{command
73 (HOL) "record"} (see \secref{sec:hol-record}) or @{command (HOL)
74 "datatype"} (see \secref{sec:hol-datatype}).
78 section {* Adhoc tuples *}
81 \begin{matharray}{rcl}
82 @{attribute (HOL) split_format}@{text "\<^sup>*"} & : & \isaratt \\
86 'split\_format' (((name *) + 'and') | ('(' 'complete' ')'))
92 \item [@{attribute (HOL) split_format}~@{text "p\<^sub>1 \<dots> p\<^sub>m
93 \<AND> \<dots> \<AND> q\<^sub>1 \<dots> q\<^sub>n"}] puts expressions of
94 low-level tuple types into canonical form as specified by the
95 arguments given; the @{text i}-th collection of arguments refers to
96 occurrences in premise @{text i} of the rule. The ``@{text
97 "(complete)"}'' option causes \emph{all} arguments in function
98 applications to be represented canonically according to their tuple
101 Note that these operations tend to invent funny names for new local
102 parameters to be introduced.
108 section {* Records \label{sec:hol-record} *}
111 In principle, records merely generalize the concept of tuples, where
112 components may be addressed by labels instead of just position. The
113 logical infrastructure of records in Isabelle/HOL is slightly more
114 advanced, though, supporting truly extensible record schemes. This
115 admits operations that are polymorphic with respect to record
116 extension, yielding ``object-oriented'' effects like (single)
117 inheritance. See also \cite{NaraschewskiW-TPHOLs98} for more
118 details on object-oriented verification and record subtyping in HOL.
122 subsection {* Basic concepts *}
125 Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records
126 at the level of terms and types. The notation is as follows:
129 \begin{tabular}{l|l|l}
130 & record terms & record types \\ \hline
131 fixed & @{text "\<lparr>x = a, y = b\<rparr>"} & @{text "\<lparr>x :: A, y :: B\<rparr>"} \\
132 schematic & @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} &
133 @{text "\<lparr>x :: A, y :: B, \<dots> :: M\<rparr>"} \\
137 \noindent The ASCII representation of @{text "\<lparr>x = a\<rparr>"} is @{text
140 A fixed record @{text "\<lparr>x = a, y = b\<rparr>"} has field @{text x} of value
141 @{text a} and field @{text y} of value @{text b}. The corresponding
142 type is @{text "\<lparr>x :: A, y :: B\<rparr>"}, assuming that @{text "a :: A"}
143 and @{text "b :: B"}.
145 A record scheme like @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} contains fields
146 @{text x} and @{text y} as before, but also possibly further fields
147 as indicated by the ``@{text "\<dots>"}'' notation (which is actually part
148 of the syntax). The improper field ``@{text "\<dots>"}'' of a record
149 scheme is called the \emph{more part}. Logically it is just a free
150 variable, which is occasionally referred to as ``row variable'' in
151 the literature. The more part of a record scheme may be
152 instantiated by zero or more further components. For example, the
153 previous scheme may get instantiated to @{text "\<lparr>x = a, y = b, z =
154 c, \<dots> = m'\<rparr>"}, where @{text m'} refers to a different more part.
155 Fixed records are special instances of record schemes, where
156 ``@{text "\<dots>"}'' is properly terminated by the @{text "() :: unit"}
157 element. In fact, @{text "\<lparr>x = a, y = b\<rparr>"} is just an abbreviation
158 for @{text "\<lparr>x = a, y = b, \<dots> = ()\<rparr>"}.
160 \medskip Two key observations make extensible records in a simply
161 typed language like HOL work out:
165 \item the more part is internalized, as a free term or type
168 \item field names are externalized, they cannot be accessed within
169 the logic as first-class values.
173 \medskip In Isabelle/HOL record types have to be defined explicitly,
174 fixing their field names and types, and their (optional) parent
175 record. Afterwards, records may be formed using above syntax, while
176 obeying the canonical order of fields as given by their declaration.
177 The record package provides several standard operations like
178 selectors and updates. The common setup for various generic proof
179 tools enable succinct reasoning patterns. See also the Isabelle/HOL
180 tutorial \cite{isabelle-hol-book} for further instructions on using
185 subsection {* Record specifications *}
188 \begin{matharray}{rcl}
189 @{command_def (HOL) "record"} & : & \isartrans{theory}{theory} \\
193 'record' typespec '=' (type '+')? (constdecl +)
199 \item [@{command (HOL) "record"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t
200 = \<tau> + c\<^sub>1 :: \<sigma>\<^sub>1 \<dots> c\<^sub>n :: \<sigma>\<^sub>n"}] defines
201 extensible record type @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"},
202 derived from the optional parent record @{text "\<tau>"} by adding new
203 field components @{text "c\<^sub>i :: \<sigma>\<^sub>i"} etc.
205 The type variables of @{text "\<tau>"} and @{text "\<sigma>\<^sub>i"} need to be
206 covered by the (distinct) parameters @{text "\<alpha>\<^sub>1, \<dots>,
207 \<alpha>\<^sub>m"}. Type constructor @{text t} has to be new, while @{text
208 \<tau>} needs to specify an instance of an existing record type. At
209 least one new field @{text "c\<^sub>i"} has to be specified.
210 Basically, field names need to belong to a unique record. This is
211 not a real restriction in practice, since fields are qualified by
212 the record name internally.
214 The parent record specification @{text \<tau>} is optional; if omitted
215 @{text t} becomes a root record. The hierarchy of all records
216 declared within a theory context forms a forest structure, i.e.\ a
217 set of trees starting with a root record each. There is no way to
218 merge multiple parent records!
220 For convenience, @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} is made a
221 type abbreviation for the fixed record type @{text "\<lparr>c\<^sub>1 ::
222 \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n\<rparr>"}, likewise is @{text
223 "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m, \<zeta>) t_scheme"} made an abbreviation for
224 @{text "\<lparr>c\<^sub>1 :: \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n, \<dots> ::
231 subsection {* Record operations *}
234 Any record definition of the form presented above produces certain
235 standard operations. Selectors and updates are provided for any
236 field, including the improper one ``@{text more}''. There are also
237 cumulative record constructor functions. To simplify the
238 presentation below, we assume for now that @{text "(\<alpha>\<^sub>1, \<dots>,
239 \<alpha>\<^sub>m) t"} is a root record with fields @{text "c\<^sub>1 ::
240 \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n"}.
242 \medskip \textbf{Selectors} and \textbf{updates} are available for
243 any field (including ``@{text more}''):
245 \begin{matharray}{lll}
246 @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
247 @{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>"} \\
250 There is special syntax for application of updates: @{text "r\<lparr>x :=
251 a\<rparr>"} abbreviates term @{text "x_update a r"}. Further notation for
252 repeated updates is also available: @{text "r\<lparr>x := a\<rparr>\<lparr>y := b\<rparr>\<lparr>z :=
253 c\<rparr>"} may be written @{text "r\<lparr>x := a, y := b, z := c\<rparr>"}. Note that
254 because of postfix notation the order of fields shown here is
255 reverse than in the actual term. Since repeated updates are just
256 function applications, fields may be freely permuted in @{text "\<lparr>x
257 := a, y := b, z := c\<rparr>"}, as far as logical equality is concerned.
258 Thus commutativity of independent updates can be proven within the
259 logic for any two fields, but not as a general theorem.
261 \medskip The \textbf{make} operation provides a cumulative record
262 constructor function:
264 \begin{matharray}{lll}
265 @{text "t.make"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
268 \medskip We now reconsider the case of non-root records, which are
269 derived of some parent. In general, the latter may depend on
270 another parent as well, resulting in a list of \emph{ancestor
271 records}. Appending the lists of fields of all ancestors results in
272 a certain field prefix. The record package automatically takes care
273 of this by lifting operations over this context of ancestor fields.
274 Assuming that @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} has ancestor
275 fields @{text "b\<^sub>1 :: \<rho>\<^sub>1, \<dots>, b\<^sub>k :: \<rho>\<^sub>k"},
276 the above record operations will get the following types:
280 @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
281 @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow>
282 \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow>
283 \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
284 @{text "t.make"} & @{text "::"} & @{text "\<rho>\<^sub>1 \<Rightarrow> \<dots> \<rho>\<^sub>k \<Rightarrow> \<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow>
285 \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
289 \noindent Some further operations address the extension aspect of a
290 derived record scheme specifically: @{text "t.fields"} produces a
291 record fragment consisting of exactly the new fields introduced here
292 (the result may serve as a more part elsewhere); @{text "t.extend"}
293 takes a fixed record and adds a given more part; @{text
294 "t.truncate"} restricts a record scheme to a fixed record.
298 @{text "t.fields"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
299 @{text "t.extend"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr> \<Rightarrow>
300 \<zeta> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
301 @{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>"} \\
305 \noindent Note that @{text "t.make"} and @{text "t.fields"} coincide
310 subsection {* Derived rules and proof tools *}
313 The record package proves several results internally, declaring
314 these facts to appropriate proof tools. This enables users to
315 reason about record structures quite conveniently. Assume that
316 @{text t} is a record type as specified above.
320 \item Standard conversions for selectors or updates applied to
321 record constructor terms are made part of the default Simplifier
322 context; thus proofs by reduction of basic operations merely require
323 the @{method simp} method without further arguments. These rules
324 are available as @{text "t.simps"}, too.
326 \item Selectors applied to updated records are automatically reduced
327 by an internal simplification procedure, which is also part of the
328 standard Simplifier setup.
330 \item Inject equations of a form analogous to @{prop "(x, y) = (x',
331 y') \<equiv> x = x' \<and> y = y'"} are declared to the Simplifier and Classical
332 Reasoner as @{attribute iff} rules. These rules are available as
335 \item The introduction rule for record equality analogous to @{text
336 "x r = x r' \<Longrightarrow> y r = y r' \<dots> \<Longrightarrow> r = r'"} is declared to the Simplifier,
337 and as the basic rule context as ``@{attribute intro}@{text "?"}''.
338 The rule is called @{text "t.equality"}.
340 \item Representations of arbitrary record expressions as canonical
341 constructor terms are provided both in @{method cases} and @{method
342 induct} format (cf.\ the generic proof methods of the same name,
343 \secref{sec:cases-induct}). Several variations are available, for
344 fixed records, record schemes, more parts etc.
346 The generic proof methods are sufficiently smart to pick the most
347 sensible rule according to the type of the indicated record
348 expression: users just need to apply something like ``@{text "(cases
349 r)"}'' to a certain proof problem.
351 \item The derived record operations @{text "t.make"}, @{text
352 "t.fields"}, @{text "t.extend"}, @{text "t.truncate"} are \emph{not}
353 treated automatically, but usually need to be expanded by hand,
354 using the collective fact @{text "t.defs"}.
360 section {* Datatypes \label{sec:hol-datatype} *}
363 \begin{matharray}{rcl}
364 @{command_def (HOL) "datatype"} & : & \isartrans{theory}{theory} \\
365 @{command_def (HOL) "rep_datatype"} & : & \isartrans{theory}{theory} \\
369 'datatype' (dtspec + 'and')
371 'rep\_datatype' (name *) dtrules
374 dtspec: parname? typespec infix? '=' (cons + '|')
376 cons: name (type *) mixfix?
378 dtrules: 'distinct' thmrefs 'inject' thmrefs 'induction' thmrefs
383 \item [@{command (HOL) "datatype"}] defines inductive datatypes in
386 \item [@{command (HOL) "rep_datatype"}] represents existing types as
387 inductive ones, generating the standard infrastructure of derived
388 concepts (primitive recursion etc.).
392 The induction and exhaustion theorems generated provide case names
393 according to the constructors involved, while parameters are named
394 after the types (see also \secref{sec:cases-induct}).
396 See \cite{isabelle-HOL} for more details on datatypes, but beware of
397 the old-style theory syntax being used there! Apart from proper
398 proof methods for case-analysis and induction, there are also
399 emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL)
400 induct_tac} available, see \secref{sec:hol-induct-tac}; these admit
401 to refer directly to the internal structure of subgoals (including
402 internally bound parameters).
406 section {* Recursive functions \label{sec:recursion} *}
409 \begin{matharray}{rcl}
410 @{command_def (HOL) "primrec"} & : & \isarkeep{local{\dsh}theory} \\
411 @{command_def (HOL) "fun"} & : & \isarkeep{local{\dsh}theory} \\
412 @{command_def (HOL) "function"} & : & \isartrans{local{\dsh}theory}{proof(prove)} \\
413 @{command_def (HOL) "termination"} & : & \isartrans{local{\dsh}theory}{proof(prove)} \\
417 'primrec' target? fixes 'where' equations
419 equations: (thmdecl? prop + '|')
421 ('fun' | 'function') target? functionopts? fixes 'where' clauses
423 clauses: (thmdecl? prop ('(' 'otherwise' ')')? + '|')
425 functionopts: '(' (('sequential' | 'domintros' | 'tailrec' | 'default' term) + ',') ')'
427 'termination' ( term )?
432 \item [@{command (HOL) "primrec"}] defines primitive recursive
433 functions over datatypes, see also \cite{isabelle-HOL}.
435 \item [@{command (HOL) "function"}] defines functions by general
436 wellfounded recursion. A detailed description with examples can be
437 found in \cite{isabelle-function}. The function is specified by a
438 set of (possibly conditional) recursive equations with arbitrary
439 pattern matching. The command generates proof obligations for the
440 completeness and the compatibility of patterns.
442 The defined function is considered partial, and the resulting
443 simplification rules (named @{text "f.psimps"}) and induction rule
444 (named @{text "f.pinduct"}) are guarded by a generated domain
445 predicate @{text "f_dom"}. The @{command (HOL) "termination"}
446 command can then be used to establish that the function is total.
448 \item [@{command (HOL) "fun"}] is a shorthand notation for
449 ``@{command (HOL) "function"}~@{text "(sequential)"}, followed by
450 automated proof attempts regarding pattern matching and termination.
451 See \cite{isabelle-function} for further details.
453 \item [@{command (HOL) "termination"}~@{text f}] commences a
454 termination proof for the previously defined function @{text f}. If
455 this is omitted, the command refers to the most recent function
456 definition. After the proof is closed, the recursive equations and
457 the induction principle is established.
463 Recursive definitions introduced by both the @{command (HOL)
464 "primrec"} and the @{command (HOL) "function"} command accommodate
465 reasoning by induction (cf.\ \secref{sec:cases-induct}): rule @{text
466 "c.induct"} (where @{text c} is the name of the function definition)
467 refers to a specific induction rule, with parameters named according
468 to the user-specified equations. Case names of @{command (HOL)
469 "primrec"} are that of the datatypes involved, while those of
470 @{command (HOL) "function"} are numbered (starting from 1).
472 The equations provided by these packages may be referred later as
473 theorem list @{text "f.simps"}, where @{text f} is the (collective)
474 name of the functions defined. Individual equations may be named
477 The @{command (HOL) "function"} command accepts the following
482 \item [@{text sequential}] enables a preprocessor which
483 disambiguates overlapping patterns by making them mutually disjoint.
484 Earlier equations take precedence over later ones. This allows to
485 give the specification in a format very similar to functional
486 programming. Note that the resulting simplification and induction
487 rules correspond to the transformed specification, not the one given
488 originally. This usually means that each equation given by the user
489 may result in several theroems. Also note that this automatic
490 transformation only works for ML-style datatype patterns.
492 \item [@{text domintros}] enables the automated generation of
493 introduction rules for the domain predicate. While mostly not
494 needed, they can be helpful in some proofs about partial functions.
496 \item [@{text tailrec}] generates the unconstrained recursive
497 equations even without a termination proof, provided that the
498 function is tail-recursive. This currently only works
500 \item [@{text "default d"}] allows to specify a default value for a
501 (partial) function, which will ensure that @{text "f x = d x"}
502 whenever @{text "x \<notin> f_dom"}.
508 subsection {* Proof methods related to recursive definitions *}
511 \begin{matharray}{rcl}
512 @{method_def (HOL) pat_completeness} & : & \isarmeth \\
513 @{method_def (HOL) relation} & : & \isarmeth \\
514 @{method_def (HOL) lexicographic_order} & : & \isarmeth \\
520 'lexicographic\_order' (clasimpmod *)
526 \item [@{method (HOL) pat_completeness}] is a specialized method to
527 solve goals regarding the completeness of pattern matching, as
528 required by the @{command (HOL) "function"} package (cf.\
529 \cite{isabelle-function}).
531 \item [@{method (HOL) relation}~@{text R}] introduces a termination
532 proof using the relation @{text R}. The resulting proof state will
533 contain goals expressing that @{text R} is wellfounded, and that the
534 arguments of recursive calls decrease with respect to @{text R}.
535 Usually, this method is used as the initial proof step of manual
538 \item [@{method (HOL) "lexicographic_order"}] attempts a fully
539 automated termination proof by searching for a lexicographic
540 combination of size measures on the arguments of the function. The
541 method accepts the same arguments as the @{method auto} method,
542 which it uses internally to prove local descents. The same context
543 modifiers as for @{method auto} are accepted, see
544 \secref{sec:clasimp}.
546 In case of failure, extensive information is printed, which can help
547 to analyse the situation (cf.\ \cite{isabelle-function}).
553 subsection {* Old-style recursive function definitions (TFL) *}
556 The old TFL commands @{command (HOL) "recdef"} and @{command (HOL)
557 "recdef_tc"} for defining recursive are mostly obsolete; @{command
558 (HOL) "function"} or @{command (HOL) "fun"} should be used instead.
560 \begin{matharray}{rcl}
561 @{command_def (HOL) "recdef"} & : & \isartrans{theory}{theory} \\
562 @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & \isartrans{theory}{proof(prove)} \\
566 'recdef' ('(' 'permissive' ')')? \\ name term (prop +) hints?
570 hints: '(' 'hints' (recdefmod *) ')'
572 recdefmod: (('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del') ':' thmrefs) | clasimpmod
574 tc: nameref ('(' nat ')')?
580 \item [@{command (HOL) "recdef"}] defines general well-founded
581 recursive functions (using the TFL package), see also
582 \cite{isabelle-HOL}. The ``@{text "(permissive)"}'' option tells
583 TFL to recover from failed proof attempts, returning unfinished
584 results. The @{text recdef_simp}, @{text recdef_cong}, and @{text
585 recdef_wf} hints refer to auxiliary rules to be used in the internal
586 automated proof process of TFL. Additional @{syntax clasimpmod}
587 declarations (cf.\ \secref{sec:clasimp}) may be given to tune the
588 context of the Simplifier (cf.\ \secref{sec:simplifier}) and
589 Classical reasoner (cf.\ \secref{sec:classical}).
591 \item [@{command (HOL) "recdef_tc"}~@{text "c (i)"}] recommences the
592 proof for leftover termination condition number @{text i} (default
593 1) as generated by a @{command (HOL) "recdef"} definition of
596 Note that in most cases, @{command (HOL) "recdef"} is able to finish
597 its internal proofs without manual intervention.
601 \medskip Hints for @{command (HOL) "recdef"} may be also declared
602 globally, using the following attributes.
604 \begin{matharray}{rcl}
605 @{attribute_def (HOL) recdef_simp} & : & \isaratt \\
606 @{attribute_def (HOL) recdef_cong} & : & \isaratt \\
607 @{attribute_def (HOL) recdef_wf} & : & \isaratt \\
611 ('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del')
617 section {* Definition by specification \label{sec:hol-specification} *}
620 \begin{matharray}{rcl}
621 @{command_def (HOL) "specification"} & : & \isartrans{theory}{proof(prove)} \\
622 @{command_def (HOL) "ax_specification"} & : & \isartrans{theory}{proof(prove)} \\
626 ('specification' | 'ax\_specification') '(' (decl +) ')' \\ (thmdecl? prop +)
628 decl: ((name ':')? term '(' 'overloaded' ')'?)
633 \item [@{command (HOL) "specification"}~@{text "decls \<phi>"}] sets up a
634 goal stating the existence of terms with the properties specified to
635 hold for the constants given in @{text decls}. After finishing the
636 proof, the theory will be augmented with definitions for the given
637 constants, as well as with theorems stating the properties for these
640 \item [@{command (HOL) "ax_specification"}~@{text "decls \<phi>"}] sets
641 up a goal stating the existence of terms with the properties
642 specified to hold for the constants given in @{text decls}. After
643 finishing the proof, the theory will be augmented with axioms
644 expressing the properties given in the first place.
646 \item [@{text decl}] declares a constant to be defined by the
647 specification given. The definition for the constant @{text c} is
648 bound to the name @{text c_def} unless a theorem name is given in
649 the declaration. Overloaded constants should be declared as such.
653 Whether to use @{command (HOL) "specification"} or @{command (HOL)
654 "ax_specification"} is to some extent a matter of style. @{command
655 (HOL) "specification"} introduces no new axioms, and so by
656 construction cannot introduce inconsistencies, whereas @{command
657 (HOL) "ax_specification"} does introduce axioms, but only after the
658 user has explicitly proven it to be safe. A practical issue must be
659 considered, though: After introducing two constants with the same
660 properties using @{command (HOL) "specification"}, one can prove
661 that the two constants are, in fact, equal. If this might be a
662 problem, one should use @{command (HOL) "ax_specification"}.
666 section {* Inductive and coinductive definitions \label{sec:hol-inductive} *}
669 An \textbf{inductive definition} specifies the least predicate (or
670 set) @{text R} closed under given rules: applying a rule to elements
671 of @{text R} yields a result within @{text R}. For example, a
672 structural operational semantics is an inductive definition of an
675 Dually, a \textbf{coinductive definition} specifies the greatest
676 predicate~/ set @{text R} that is consistent with given rules: every
677 element of @{text R} can be seen as arising by applying a rule to
678 elements of @{text R}. An important example is using bisimulation
679 relations to formalise equivalence of processes and infinite data
682 \medskip The HOL package is related to the ZF one, which is
683 described in a separate paper,\footnote{It appeared in CADE
684 \cite{paulson-CADE}; a longer version is distributed with Isabelle.}
685 which you should refer to in case of difficulties. The package is
686 simpler than that of ZF thanks to implicit type-checking in HOL.
687 The types of the (co)inductive predicates (or sets) determine the
688 domain of the fixedpoint definition, and the package does not have
689 to use inference rules for type-checking.
691 \begin{matharray}{rcl}
692 @{command_def (HOL) "inductive"} & : & \isarkeep{local{\dsh}theory} \\
693 @{command_def (HOL) "inductive_set"} & : & \isarkeep{local{\dsh}theory} \\
694 @{command_def (HOL) "coinductive"} & : & \isarkeep{local{\dsh}theory} \\
695 @{command_def (HOL) "coinductive_set"} & : & \isarkeep{local{\dsh}theory} \\
696 @{attribute_def (HOL) mono} & : & \isaratt \\
700 ('inductive' | 'inductive\_set' | 'coinductive' | 'coinductive\_set') target? fixes ('for' fixes)? \\
701 ('where' clauses)? ('monos' thmrefs)?
703 clauses: (thmdecl? prop + '|')
705 'mono' (() | 'add' | 'del')
711 \item [@{command (HOL) "inductive"} and @{command (HOL)
712 "coinductive"}] define (co)inductive predicates from the
713 introduction rules given in the @{keyword "where"} part. The
714 optional @{keyword "for"} part contains a list of parameters of the
715 (co)inductive predicates that remain fixed throughout the
716 definition. The optional @{keyword "monos"} section contains
717 \emph{monotonicity theorems}, which are required for each operator
718 applied to a recursive set in the introduction rules. There
719 \emph{must} be a theorem of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"},
720 for each premise @{text "M R\<^sub>i t"} in an introduction rule!
722 \item [@{command (HOL) "inductive_set"} and @{command (HOL)
723 "coinductive_set"}] are wrappers for to the previous commands,
724 allowing the definition of (co)inductive sets.
726 \item [@{attribute (HOL) mono}] declares monotonicity rules. These
727 rule are involved in the automated monotonicity proof of @{command
734 subsection {* Derived rules *}
737 Each (co)inductive definition @{text R} adds definitions to the
738 theory and also proves some theorems:
742 \item [@{text R.intros}] is the list of introduction rules as proven
743 theorems, for the recursive predicates (or sets). The rules are
744 also available individually, using the names given them in the
747 \item [@{text R.cases}] is the case analysis (or elimination) rule;
749 \item [@{text R.induct} or @{text R.coinduct}] is the (co)induction
754 When several predicates @{text "R\<^sub>1, \<dots>, R\<^sub>n"} are
755 defined simultaneously, the list of introduction rules is called
756 @{text "R\<^sub>1_\<dots>_R\<^sub>n.intros"}, the case analysis rules are
757 called @{text "R\<^sub>1.cases, \<dots>, R\<^sub>n.cases"}, and the list
758 of mutual induction rules is called @{text
759 "R\<^sub>1_\<dots>_R\<^sub>n.inducts"}.
763 subsection {* Monotonicity theorems *}
766 Each theory contains a default set of theorems that are used in
767 monotonicity proofs. New rules can be added to this set via the
768 @{attribute (HOL) mono} attribute. The HOL theory @{text Inductive}
769 shows how this is done. In general, the following monotonicity
770 theorems may be added:
774 \item Theorems of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"}, for proving
775 monotonicity of inductive definitions whose introduction rules have
776 premises involving terms such as @{text "M R\<^sub>i t"}.
778 \item Monotonicity theorems for logical operators, which are of the
779 general form @{text "(\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> (\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> \<longrightarrow> \<dots>"}. For example, in
780 the case of the operator @{text "\<or>"}, the corresponding theorem is
782 \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"}}
785 \item De Morgan style equations for reasoning about the ``polarity''
788 @{prop "\<not> \<not> P \<longleftrightarrow> P"} \qquad\qquad
789 @{prop "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"}
792 \item Equations for reducing complex operators to more primitive
793 ones whose monotonicity can easily be proved, e.g.
795 @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> \<not> P \<or> Q"} \qquad\qquad
796 @{prop "Ball A P \<equiv> \<forall>x. x \<in> A \<longrightarrow> P x"}
801 %FIXME: Example of an inductive definition
805 section {* Arithmetic proof support *}
808 \begin{matharray}{rcl}
809 @{method_def (HOL) arith} & : & \isarmeth \\
810 @{attribute_def (HOL) arith_split} & : & \isaratt \\
813 The @{method (HOL) arith} method decides linear arithmetic problems
814 (on types @{text nat}, @{text int}, @{text real}). Any current
815 facts are inserted into the goal before running the procedure.
817 The @{attribute (HOL) arith_split} attribute declares case split
818 rules to be expanded before the arithmetic procedure is invoked.
820 Note that a simpler (but faster) version of arithmetic reasoning is
821 already performed by the Simplifier.
825 section {* Cases and induction: emulating tactic scripts \label{sec:hol-induct-tac} *}
828 The following important tactical tools of Isabelle/HOL have been
829 ported to Isar. These should be never used in proper proof texts!
831 \begin{matharray}{rcl}
832 @{method_def (HOL) case_tac}@{text "\<^sup>*"} & : & \isarmeth \\
833 @{method_def (HOL) induct_tac}@{text "\<^sup>*"} & : & \isarmeth \\
834 @{method_def (HOL) ind_cases}@{text "\<^sup>*"} & : & \isarmeth \\
835 @{command_def (HOL) "inductive_cases"} & : & \isartrans{theory}{theory} \\
839 'case\_tac' goalspec? term rule?
841 'induct\_tac' goalspec? (insts * 'and') rule?
843 'ind\_cases' (prop +) ('for' (name +)) ?
845 'inductive\_cases' (thmdecl? (prop +) + 'and')
848 rule: ('rule' ':' thmref)
854 \item [@{method (HOL) case_tac} and @{method (HOL) induct_tac}]
855 admit to reason about inductive datatypes only (unless an
856 alternative rule is given explicitly). Furthermore, @{method (HOL)
857 case_tac} does a classical case split on booleans; @{method (HOL)
858 induct_tac} allows only variables to be given as instantiation.
859 These tactic emulations feature both goal addressing and dynamic
860 instantiation. Note that named rule cases are \emph{not} provided
861 as would be by the proper @{method induct} and @{method cases} proof
862 methods (see \secref{sec:cases-induct}).
864 \item [@{method (HOL) ind_cases} and @{command (HOL)
865 "inductive_cases"}] provide an interface to the internal @{ML_text
866 mk_cases} operation. Rules are simplified in an unrestricted
869 While @{method (HOL) ind_cases} is a proof method to apply the
870 result immediately as elimination rules, @{command (HOL)
871 "inductive_cases"} provides case split theorems at the theory level
872 for later use. The @{keyword "for"} argument of the @{method (HOL)
873 ind_cases} method allows to specify a list of variables that should
874 be generalized before applying the resulting rule.
880 section {* Executable code *}
883 Isabelle/Pure provides two generic frameworks to support code
884 generation from executable specifications. Isabelle/HOL
885 instantiates these mechanisms in a way that is amenable to end-user
888 One framework generates code from both functional and relational
889 programs to SML. See \cite{isabelle-HOL} for further information
890 (this actually covers the new-style theory format as well).
892 \begin{matharray}{rcl}
893 @{command_def (HOL) "value"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
894 @{command_def (HOL) "code_module"} & : & \isartrans{theory}{theory} \\
895 @{command_def (HOL) "code_library"} & : & \isartrans{theory}{theory} \\
896 @{command_def (HOL) "consts_code"} & : & \isartrans{theory}{theory} \\
897 @{command_def (HOL) "types_code"} & : & \isartrans{theory}{theory} \\
898 @{attribute_def (HOL) code} & : & \isaratt \\
905 ( 'code\_module' | 'code\_library' ) modespec ? name ? \\
906 ( 'file' name ) ? ( 'imports' ( name + ) ) ? \\
907 'contains' ( ( name '=' term ) + | term + )
910 modespec: '(' ( name * ) ')'
913 'consts\_code' (codespec +)
916 codespec: const template attachment ?
919 'types\_code' (tycodespec +)
922 tycodespec: name template attachment ?
928 template: '(' string ')'
931 attachment: 'attach' modespec ? verblbrace text verbrbrace
940 \item [@{command (HOL) "value"}~@{text t}] evaluates and prints a
941 term using the code generator.
945 \medskip The other framework generates code from functional programs
946 (including overloading using type classes) to SML \cite{SML}, OCaml
947 \cite{OCaml} and Haskell \cite{haskell-revised-report}.
948 Conceptually, code generation is split up in three steps:
949 \emph{selection} of code theorems, \emph{translation} into an
950 abstract executable view and \emph{serialization} to a specific
951 \emph{target language}. See \cite{isabelle-codegen} for an
952 introduction on how to use it.
954 \begin{matharray}{rcl}
955 @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
956 @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
957 @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
958 @{command_def (HOL) "code_datatype"} & : & \isartrans{theory}{theory} \\
959 @{command_def (HOL) "code_const"} & : & \isartrans{theory}{theory} \\
960 @{command_def (HOL) "code_type"} & : & \isartrans{theory}{theory} \\
961 @{command_def (HOL) "code_class"} & : & \isartrans{theory}{theory} \\
962 @{command_def (HOL) "code_instance"} & : & \isartrans{theory}{theory} \\
963 @{command_def (HOL) "code_monad"} & : & \isartrans{theory}{theory} \\
964 @{command_def (HOL) "code_reserved"} & : & \isartrans{theory}{theory} \\
965 @{command_def (HOL) "code_include"} & : & \isartrans{theory}{theory} \\
966 @{command_def (HOL) "code_modulename"} & : & \isartrans{theory}{theory} \\
967 @{command_def (HOL) "code_exception"} & : & \isartrans{theory}{theory} \\
968 @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
969 @{attribute_def (HOL) code} & : & \isaratt \\
973 'export\_code' ( constexpr + ) ? \\
974 ( ( 'in' target ( 'module\_name' string ) ? \\
975 ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
978 'code\_thms' ( constexpr + ) ?
981 'code\_deps' ( constexpr + ) ?
987 constexpr: ( const | 'name.*' | '*' )
990 typeconstructor: nameref
996 target: 'OCaml' | 'SML' | 'Haskell'
999 'code\_datatype' const +
1002 'code\_const' (const + 'and') \\
1003 ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
1006 'code\_type' (typeconstructor + 'and') \\
1007 ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
1010 'code\_class' (class + 'and') \\
1012 ( ( string ('where' \\
1013 ( const ( '==' | equiv ) string ) + ) ? ) ? + 'and' ) ')' ) + )
1016 'code\_instance' (( typeconstructor '::' class ) + 'and') \\
1017 ( ( '(' target ( '-' ? + 'and' ) ')' ) + )
1020 'code\_monad' const const target
1023 'code\_reserved' target ( string + )
1026 'code\_include' target ( string ( string | '-') )
1029 'code\_modulename' target ( ( string string ) + )
1032 'code\_exception' ( const + )
1035 syntax: string | ( 'infix' | 'infixl' | 'infixr' ) nat string
1038 'code' ('func' | 'inline') ( 'del' )?
1044 \item [@{command (HOL) "export_code"}] is the canonical interface
1045 for generating and serializing code: for a given list of constants,
1046 code is generated for the specified target languages. Abstract code
1047 is cached incrementally. If no constant is given, the currently
1048 cached code is serialized. If no serialization instruction is
1049 given, only abstract code is cached.
1051 Constants may be specified by giving them literally, referring to
1052 all executable contants within a certain theory by giving @{text
1053 "name.*"}, or referring to \emph{all} executable constants currently
1054 available by giving @{text "*"}.
1056 By default, for each involved theory one corresponding name space
1057 module is generated. Alternativly, a module name may be specified
1058 after the @{keyword "module_name"} keyword; then \emph{all} code is
1059 placed in this module.
1061 For \emph{SML} and \emph{OCaml}, the file specification refers to a
1062 single file; for \emph{Haskell}, it refers to a whole directory,
1063 where code is generated in multiple files reflecting the module
1064 hierarchy. The file specification ``@{text "-"}'' denotes standard
1065 output. For \emph{SML}, omitting the file specification compiles
1066 code internally in the context of the current ML session.
1068 Serializers take an optional list of arguments in parentheses. For
1069 \emph{Haskell} a module name prefix may be given using the ``@{text
1070 "root:"}'' argument; ``@{text string_classes}'' adds a ``@{verbatim
1071 "deriving (Read, Show)"}'' clause to each appropriate datatype
1074 \item [@{command (HOL) "code_thms"}] prints a list of theorems
1075 representing the corresponding program containing all given
1076 constants; if no constants are given, the currently cached code
1077 theorems are printed.
1079 \item [@{command (HOL) "code_deps"}] visualizes dependencies of
1080 theorems representing the corresponding program containing all given
1081 constants; if no constants are given, the currently cached code
1082 theorems are visualized.
1084 \item [@{command (HOL) "code_datatype"}] specifies a constructor set
1087 \item [@{command (HOL) "code_const"}] associates a list of constants
1088 with target-specific serializations; omitting a serialization
1089 deletes an existing serialization.
1091 \item [@{command (HOL) "code_type"}] associates a list of type
1092 constructors with target-specific serializations; omitting a
1093 serialization deletes an existing serialization.
1095 \item [@{command (HOL) "code_class"}] associates a list of classes
1096 with target-specific class names; in addition, constants associated
1097 with this class may be given target-specific names used for instance
1098 declarations; omitting a serialization deletes an existing
1099 serialization. This applies only to \emph{Haskell}.
1101 \item [@{command (HOL) "code_instance"}] declares a list of type
1102 constructor / class instance relations as ``already present'' for a
1103 given target. Omitting a ``@{text "-"}'' deletes an existing
1104 ``already present'' declaration. This applies only to
1107 \item [@{command (HOL) "code_monad"}] provides an auxiliary
1108 mechanism to generate monadic code.
1110 \item [@{command (HOL) "code_reserved"}] declares a list of names as
1111 reserved for a given target, preventing it to be shadowed by any
1114 \item [@{command (HOL) "code_include"}] adds arbitrary named content
1115 (``include'') to generated code. A as last argument ``@{text "-"}''
1116 will remove an already added ``include''.
1118 \item [@{command (HOL) "code_modulename"}] declares aliasings from
1119 one module name onto another.
1121 \item [@{command (HOL) "code_exception"}] declares constants which
1122 are not required to have a definition by a defining equations; these
1123 are mapped on exceptions instead.
1125 \item [@{attribute (HOL) code}~@{text func}] explicitly selects (or
1126 with option ``@{text "del:"}'' deselects) a defining equation for
1127 code generation. Usually packages introducing defining equations
1128 provide a resonable default setup for selection.
1130 \item [@{attribute (HOL) code}@{text inline}] declares (or with
1131 option ``@{text "del:"}'' removes) inlining theorems which are
1132 applied as rewrite rules to any defining equation during
1135 \item [@{command (HOL) "print_codesetup"}] gives an overview on
1136 selected defining equations, code generator datatypes and