doc-src/IsarRef/Thy/HOL_Specific.thy
author wenzelm
Thu, 08 May 2008 22:05:15 +0200
changeset 26852 a31203f58b20
parent 26849 df50bc1249d7
child 26860 7c749112261c
permissions -rw-r--r--
misc tuning;
     1 (* $Id$ *)
     2 
     3 theory HOL_Specific
     4 imports Main
     5 begin
     6 
     7 chapter {* Isabelle/HOL \label{ch:hol} *}
     8 
     9 section {* Primitive types \label{sec:hol-typedef} *}
    10 
    11 text {*
    12   \begin{matharray}{rcl}
    13     @{command_def (HOL) "typedecl"} & : & \isartrans{theory}{theory} \\
    14     @{command_def (HOL) "typedef"} & : & \isartrans{theory}{proof(prove)} \\
    15   \end{matharray}
    16 
    17   \begin{rail}
    18     'typedecl' typespec infix?
    19     ;
    20     'typedef' altname? abstype '=' repset
    21     ;
    22 
    23     altname: '(' (name | 'open' | 'open' name) ')'
    24     ;
    25     abstype: typespec infix?
    26     ;
    27     repset: term ('morphisms' name name)?
    28     ;
    29   \end{rail}
    30 
    31   \begin{descr}
    32   
    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
    38   
    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}.
    44   
    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).
    50   
    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 @{method simp} or @{method iff} declarations).
    57   Rules @{text Rep_t_cases}/@{text Rep_t_induct}, and @{text
    58   Abs_t_cases}/@{text Abs_t_induct} provide alternative views on
    59   surjectivity; these are already declared as set or type rules for
    60   the generic @{method cases} and @{method induct} methods.
    61   
    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
    65   representing set.
    66 
    67   \end{descr}
    68 
    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}).
    75 *}
    76 
    77 
    78 section {* Adhoc tuples *}
    79 
    80 text {*
    81   \begin{matharray}{rcl}
    82     @{attribute (HOL) split_format}@{text "\<^sup>*"} & : & \isaratt \\
    83   \end{matharray}
    84 
    85   \begin{rail}
    86     'split\_format' (((name *) + 'and') | ('(' 'complete' ')'))
    87     ;
    88   \end{rail}
    89 
    90   \begin{descr}
    91   
    92   \item [@{method (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
    99   type structure.
   100 
   101   Note that these operations tend to invent funny names for new local
   102   parameters to be introduced.
   103 
   104   \end{descr}
   105 *}
   106 
   107 
   108 section {* Records \label{sec:hol-record} *}
   109 
   110 text {*
   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.
   119 *}
   120 
   121 
   122 subsection {* Basic concepts *}
   123 
   124 text {*
   125   Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records
   126   at the level of terms and types.  The notation is as follows:
   127 
   128   \begin{center}
   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>"} \\
   134   \end{tabular}
   135   \end{center}
   136 
   137   \noindent The ASCII representation of @{text "\<lparr>x = a\<rparr>"} is @{text
   138   "(| x = a |)"}.
   139 
   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"}.
   144 
   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>"}.
   159   
   160   \medskip Two key observations make extensible records in a simply
   161   typed language like HOL work out:
   162 
   163   \begin{enumerate}
   164 
   165   \item the more part is internalized, as a free term or type
   166   variable,
   167 
   168   \item field names are externalized, they cannot be accessed within
   169   the logic as first-class values.
   170 
   171   \end{enumerate}
   172 
   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
   181   records in practice.
   182 *}
   183 
   184 
   185 subsection {* Record specifications *}
   186 
   187 text {*
   188   \begin{matharray}{rcl}
   189     @{command_def (HOL) "record"} & : & \isartrans{theory}{theory} \\
   190   \end{matharray}
   191 
   192   \begin{rail}
   193     'record' typespec '=' (type '+')? (constdecl +)
   194     ;
   195   \end{rail}
   196 
   197   \begin{descr}
   198 
   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.
   204 
   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.
   213 
   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!
   219 
   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> ::
   225   \<zeta>\<rparr>"}.
   226 
   227   \end{descr}
   228 *}
   229 
   230 
   231 subsection {* Record operations *}
   232 
   233 text {*
   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"}.
   241 
   242   \medskip \textbf{Selectors} and \textbf{updates} are available for
   243   any field (including ``@{text more}''):
   244 
   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>"} \\
   248   \end{matharray}
   249 
   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.
   260 
   261   \medskip The \textbf{make} operation provides a cumulative record
   262   constructor function:
   263 
   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>"} \\
   266   \end{matharray}
   267 
   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:
   277 
   278   \medskip
   279   \begin{tabular}{lll}
   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>"} \\
   286   \end{tabular}
   287   \medskip
   288 
   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.
   295 
   296   \medskip
   297   \begin{tabular}{lll}
   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>"} \\
   302   \end{tabular}
   303   \medskip
   304 
   305   \noindent Note that @{text "t.make"} and @{text "t.fields"} coincide
   306   for root records.
   307 *}
   308 
   309 
   310 subsection {* Derived rules and proof tools *}
   311 
   312 text {*
   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.
   317 
   318   \begin{enumerate}
   319   
   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.
   325   
   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.
   329 
   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
   333   @{text "t.iffs"}.
   334 
   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"}.
   339 
   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.
   345   
   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.
   350 
   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"}.
   355 
   356   \end{enumerate}
   357 *}
   358 
   359 
   360 section {* Datatypes \label{sec:hol-datatype} *}
   361 
   362 text {*
   363   \begin{matharray}{rcl}
   364     @{command_def (HOL) "datatype"} & : & \isartrans{theory}{theory} \\
   365     @{command_def (HOL) "rep_datatype"} & : & \isartrans{theory}{theory} \\
   366   \end{matharray}
   367 
   368   \begin{rail}
   369     'datatype' (dtspec + 'and')
   370     ;
   371     'rep\_datatype' (name *) dtrules
   372     ;
   373 
   374     dtspec: parname? typespec infix? '=' (cons + '|')
   375     ;
   376     cons: name (type *) mixfix?
   377     ;
   378     dtrules: 'distinct' thmrefs 'inject' thmrefs 'induction' thmrefs
   379   \end{rail}
   380 
   381   \begin{descr}
   382 
   383   \item [@{command (HOL) "datatype"}] defines inductive datatypes in
   384   HOL.
   385 
   386   \item [@{command (HOL) "rep_datatype"}] represents existing types as
   387   inductive ones, generating the standard infrastructure of derived
   388   concepts (primitive recursion etc.).
   389 
   390   \end{descr}
   391 
   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}).
   395 
   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).
   403 *}
   404 
   405 
   406 section {* Recursive functions \label{sec:recursion} *}
   407 
   408 text {*
   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)} \\
   414   \end{matharray}
   415 
   416   \railalias{funopts}{function\_opts}  %FIXME ??
   417 
   418   \begin{rail}
   419     'primrec' target? fixes 'where' equations
   420     ;
   421     equations: (thmdecl? prop + '|')
   422     ;
   423     ('fun' | 'function') (funopts)? fixes 'where' clauses
   424     ;
   425     clauses: (thmdecl? prop ('(' 'otherwise' ')')? + '|')
   426     ;
   427     funopts: '(' (('sequential' | 'in' name | 'domintros' | 'tailrec' |
   428     'default' term) + ',') ')'
   429     ;
   430     'termination' ( term )?
   431   \end{rail}
   432 
   433   \begin{descr}
   434 
   435   \item [@{command (HOL) "primrec"}] defines primitive recursive
   436   functions over datatypes, see also \cite{isabelle-HOL}.
   437 
   438   \item [@{command (HOL) "function"}] defines functions by general
   439   wellfounded recursion. A detailed description with examples can be
   440   found in \cite{isabelle-function}. The function is specified by a
   441   set of (possibly conditional) recursive equations with arbitrary
   442   pattern matching. The command generates proof obligations for the
   443   completeness and the compatibility of patterns.
   444 
   445   The defined function is considered partial, and the resulting
   446   simplification rules (named @{text "f.psimps"}) and induction rule
   447   (named @{text "f.pinduct"}) are guarded by a generated domain
   448   predicate @{text "f_dom"}. The @{command (HOL) "termination"}
   449   command can then be used to establish that the function is total.
   450 
   451   \item [@{command (HOL) "fun"}] is a shorthand notation for
   452   ``@{command (HOL) "function"}~@{text "(sequential)"}, followed by
   453   automated proof attempts regarding pattern matching and termination.
   454   See \cite{isabelle-function} for further details.
   455 
   456   \item [@{command (HOL) "termination"}~@{text f}] commences a
   457   termination proof for the previously defined function @{text f}.  If
   458   this is omitted, the command refers to the most recent function
   459   definition.  After the proof is closed, the recursive equations and
   460   the induction principle is established.
   461 
   462   \end{descr}
   463 
   464   %FIXME check
   465 
   466   Recursive definitions introduced by both the @{command (HOL)
   467   "primrec"} and the @{command (HOL) "function"} command accommodate
   468   reasoning by induction (cf.\ \secref{sec:cases-induct}): rule @{text
   469   "c.induct"} (where @{text c} is the name of the function definition)
   470   refers to a specific induction rule, with parameters named according
   471   to the user-specified equations.  Case names of @{command (HOL)
   472   "primrec"} are that of the datatypes involved, while those of
   473   @{command (HOL) "function"} are numbered (starting from 1).
   474 
   475   The equations provided by these packages may be referred later as
   476   theorem list @{text "f.simps"}, where @{text f} is the (collective)
   477   name of the functions defined.  Individual equations may be named
   478   explicitly as well.
   479 
   480   The @{command (HOL) "function"} command accepts the following
   481   options.
   482 
   483   \begin{descr}
   484 
   485   \item [@{text sequential}] enables a preprocessor which
   486   disambiguates overlapping patterns by making them mutually disjoint.
   487   Earlier equations take precedence over later ones.  This allows to
   488   give the specification in a format very similar to functional
   489   programming.  Note that the resulting simplification and induction
   490   rules correspond to the transformed specification, not the one given
   491   originally. This usually means that each equation given by the user
   492   may result in several theroems.  Also note that this automatic
   493   transformation only works for ML-style datatype patterns.
   494 
   495   \item [@{text "\<IN> name"}] gives the target for the definition.
   496   %FIXME ?!?
   497 
   498   \item [@{text domintros}] enables the automated generation of
   499   introduction rules for the domain predicate. While mostly not
   500   needed, they can be helpful in some proofs about partial functions.
   501 
   502   \item [@{text tailrec}] generates the unconstrained recursive
   503   equations even without a termination proof, provided that the
   504   function is tail-recursive. This currently only works
   505 
   506   \item [@{text "default d"}] allows to specify a default value for a
   507   (partial) function, which will ensure that @{text "f x = d x"}
   508   whenever @{text "x \<notin> f_dom"}.
   509 
   510   \end{descr}
   511 *}
   512 
   513 
   514 subsection {* Proof methods related to recursive definitions *}
   515 
   516 text {*
   517   \begin{matharray}{rcl}
   518     @{method_def (HOL) pat_completeness} & : & \isarmeth \\
   519     @{method_def (HOL) relation} & : & \isarmeth \\
   520     @{method_def (HOL) lexicographic_order} & : & \isarmeth \\
   521   \end{matharray}
   522 
   523   \begin{rail}
   524     'relation' term
   525     ;
   526     'lexicographic\_order' (clasimpmod *)
   527     ;
   528   \end{rail}
   529 
   530   \begin{descr}
   531 
   532   \item [@{method (HOL) pat_completeness}] is a specialized method to
   533   solve goals regarding the completeness of pattern matching, as
   534   required by the @{command (HOL) "function"} package (cf.\
   535   \cite{isabelle-function}).
   536 
   537   \item [@{method (HOL) relation}~@{text R}] introduces a termination
   538   proof using the relation @{text R}.  The resulting proof state will
   539   contain goals expressing that @{text R} is wellfounded, and that the
   540   arguments of recursive calls decrease with respect to @{text R}.
   541   Usually, this method is used as the initial proof step of manual
   542   termination proofs.
   543 
   544   \item [@{method (HOL) "lexicographic_order"}] attempts a fully
   545   automated termination proof by searching for a lexicographic
   546   combination of size measures on the arguments of the function. The
   547   method accepts the same arguments as the @{method auto} method,
   548   which it uses internally to prove local descents.  The same context
   549   modifiers as for @{method auto} are accepted, see
   550   \secref{sec:clasimp}.
   551 
   552   In case of failure, extensive information is printed, which can help
   553   to analyse the situation (cf.\ \cite{isabelle-function}).
   554 
   555   \end{descr}
   556 *}
   557 
   558 
   559 subsection {* Old-style recursive function definitions (TFL) *}
   560 
   561 text {*
   562   The old TFL commands @{command (HOL) "recdef"} and @{command (HOL)
   563   "recdef_tc"} for defining recursive are mostly obsolete; @{command
   564   (HOL) "function"} or @{command (HOL) "fun"} should be used instead.
   565 
   566   \begin{matharray}{rcl}
   567     @{command_def (HOL) "recdef"} & : & \isartrans{theory}{theory} \\
   568     @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & \isartrans{theory}{proof(prove)} \\
   569   \end{matharray}
   570 
   571   \begin{rail}
   572     'recdef' ('(' 'permissive' ')')? \\ name term (prop +) hints?
   573     ;
   574     recdeftc thmdecl? tc
   575     ;
   576     hints: '(' 'hints' (recdefmod *) ')'
   577     ;
   578     recdefmod: (('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del') ':' thmrefs) | clasimpmod
   579     ;
   580     tc: nameref ('(' nat ')')?
   581     ;
   582   \end{rail}
   583 
   584   \begin{descr}
   585   
   586   \item [@{command (HOL) "recdef"}] defines general well-founded
   587   recursive functions (using the TFL package), see also
   588   \cite{isabelle-HOL}.  The ``@{text "(permissive)"}'' option tells
   589   TFL to recover from failed proof attempts, returning unfinished
   590   results.  The @{text recdef_simp}, @{text recdef_cong}, and @{text
   591   recdef_wf} hints refer to auxiliary rules to be used in the internal
   592   automated proof process of TFL.  Additional @{syntax clasimpmod}
   593   declarations (cf.\ \secref{sec:clasimp}) may be given to tune the
   594   context of the Simplifier (cf.\ \secref{sec:simplifier}) and
   595   Classical reasoner (cf.\ \secref{sec:classical}).
   596   
   597   \item [@{command (HOL) "recdef_tc"}~@{text "c (i)"}] recommences the
   598   proof for leftover termination condition number @{text i} (default
   599   1) as generated by a @{command (HOL) "recdef"} definition of
   600   constant @{text c}.
   601   
   602   Note that in most cases, @{command (HOL) "recdef"} is able to finish
   603   its internal proofs without manual intervention.
   604 
   605   \end{descr}
   606 
   607   \medskip Hints for @{command (HOL) "recdef"} may be also declared
   608   globally, using the following attributes.
   609 
   610   \begin{matharray}{rcl}
   611     @{attribute_def (HOL) recdef_simp} & : & \isaratt \\
   612     @{attribute_def (HOL) recdef_cong} & : & \isaratt \\
   613     @{attribute_def (HOL) recdef_wf} & : & \isaratt \\
   614   \end{matharray}
   615 
   616   \begin{rail}
   617     ('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del')
   618     ;
   619   \end{rail}
   620 *}
   621 
   622 
   623 section {* Definition by specification \label{sec:hol-specification} *}
   624 
   625 text {*
   626   \begin{matharray}{rcl}
   627     @{command_def (HOL) "specification"} & : & \isartrans{theory}{proof(prove)} \\
   628     @{command_def (HOL) "ax_specification"} & : & \isartrans{theory}{proof(prove)} \\
   629   \end{matharray}
   630 
   631   \begin{rail}
   632   ('specification' | 'ax\_specification') '(' (decl +) ')' \\ (thmdecl? prop +)
   633   ;
   634   decl: ((name ':')? term '(' 'overloaded' ')'?)
   635   \end{rail}
   636 
   637   \begin{descr}
   638 
   639   \item [@{command (HOL) "specification"}~@{text "decls \<phi>"}] sets up a
   640   goal stating the existence of terms with the properties specified to
   641   hold for the constants given in @{text decls}.  After finishing the
   642   proof, the theory will be augmented with definitions for the given
   643   constants, as well as with theorems stating the properties for these
   644   constants.
   645 
   646   \item [@{command (HOL) "ax_specification"}~@{text "decls \<phi>"}] sets
   647   up a goal stating the existence of terms with the properties
   648   specified to hold for the constants given in @{text decls}.  After
   649   finishing the proof, the theory will be augmented with axioms
   650   expressing the properties given in the first place.
   651 
   652   \item [@{text decl}] declares a constant to be defined by the
   653   specification given.  The definition for the constant @{text c} is
   654   bound to the name @{text c_def} unless a theorem name is given in
   655   the declaration.  Overloaded constants should be declared as such.
   656 
   657   \end{descr}
   658 
   659   Whether to use @{command (HOL) "specification"} or @{command (HOL)
   660   "ax_specification"} is to some extent a matter of style.  @{command
   661   (HOL) "specification"} introduces no new axioms, and so by
   662   construction cannot introduce inconsistencies, whereas @{command
   663   (HOL) "ax_specification"} does introduce axioms, but only after the
   664   user has explicitly proven it to be safe.  A practical issue must be
   665   considered, though: After introducing two constants with the same
   666   properties using @{command (HOL) "specification"}, one can prove
   667   that the two constants are, in fact, equal.  If this might be a
   668   problem, one should use @{command (HOL) "ax_specification"}.
   669 *}
   670 
   671 
   672 section {* Inductive and coinductive definitions \label{sec:hol-inductive} *}
   673 
   674 text {*
   675   An \textbf{inductive definition} specifies the least predicate (or
   676   set) @{text R} closed under given rules: applying a rule to elements
   677   of @{text R} yields a result within @{text R}.  For example, a
   678   structural operational semantics is an inductive definition of an
   679   evaluation relation.
   680 
   681   Dually, a \textbf{coinductive definition} specifies the greatest
   682   predicate~/ set @{text R} that is consistent with given rules: every
   683   element of @{text R} can be seen as arising by applying a rule to
   684   elements of @{text R}.  An important example is using bisimulation
   685   relations to formalise equivalence of processes and infinite data
   686   structures.
   687 
   688   \medskip The HOL package is related to the ZF one, which is
   689   described in a separate paper,\footnote{It appeared in CADE
   690   \cite{paulson-CADE}; a longer version is distributed with Isabelle.}
   691   which you should refer to in case of difficulties.  The package is
   692   simpler than that of ZF thanks to implicit type-checking in HOL.
   693   The types of the (co)inductive predicates (or sets) determine the
   694   domain of the fixedpoint definition, and the package does not have
   695   to use inference rules for type-checking.
   696 
   697   \begin{matharray}{rcl}
   698     @{command_def (HOL) "inductive"} & : & \isarkeep{local{\dsh}theory} \\
   699     @{command_def (HOL) "inductive_set"} & : & \isarkeep{local{\dsh}theory} \\
   700     @{command_def (HOL) "coinductive"} & : & \isarkeep{local{\dsh}theory} \\
   701     @{command_def (HOL) "coinductive_set"} & : & \isarkeep{local{\dsh}theory} \\
   702     @{attribute_def (HOL) mono} & : & \isaratt \\
   703   \end{matharray}
   704 
   705   \begin{rail}
   706     ('inductive' | 'inductive\_set' | 'coinductive' | 'coinductive\_set') target? fixes ('for' fixes)? \\
   707     ('where' clauses)? ('monos' thmrefs)?
   708     ;
   709     clauses: (thmdecl? prop + '|')
   710     ;
   711     'mono' (() | 'add' | 'del')
   712     ;
   713   \end{rail}
   714 
   715   \begin{descr}
   716 
   717   \item [@{command (HOL) "inductive"} and @{command (HOL)
   718   "coinductive"}] define (co)inductive predicates from the
   719   introduction rules given in the @{keyword "where"} part.  The
   720   optional @{keyword "for"} part contains a list of parameters of the
   721   (co)inductive predicates that remain fixed throughout the
   722   definition.  The optional @{keyword "monos"} section contains
   723   \emph{monotonicity theorems}, which are required for each operator
   724   applied to a recursive set in the introduction rules.  There
   725   \emph{must} be a theorem of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"},
   726   for each premise @{text "M R\<^sub>i t"} in an introduction rule!
   727 
   728   \item [@{command (HOL) "inductive_set"} and @{command (HOL)
   729   "coinductive_set"}] are wrappers for to the previous commands,
   730   allowing the definition of (co)inductive sets.
   731 
   732   \item [@{attribute (HOL) mono}] declares monotonicity rules.  These
   733   rule are involved in the automated monotonicity proof of @{command
   734   (HOL) "inductive"}.
   735 
   736   \end{descr}
   737 *}
   738 
   739 
   740 subsection {* Derived rules *}
   741 
   742 text {*
   743   Each (co)inductive definition @{text R} adds definitions to the
   744   theory and also proves some theorems:
   745 
   746   \begin{description}
   747 
   748   \item [@{text R.intros}] is the list of introduction rules as proven
   749   theorems, for the recursive predicates (or sets).  The rules are
   750   also available individually, using the names given them in the
   751   theory file;
   752 
   753   \item [@{text R.cases}] is the case analysis (or elimination) rule;
   754 
   755   \item [@{text R.induct} or @{text R.coinduct}] is the (co)induction
   756   rule.
   757 
   758   \end{description}
   759 
   760   When several predicates @{text "R\<^sub>1, \<dots>, R\<^sub>n"} are
   761   defined simultaneously, the list of introduction rules is called
   762   @{text "R\<^sub>1_\<dots>_R\<^sub>n.intros"}, the case analysis rules are
   763   called @{text "R\<^sub>1.cases, \<dots>, R\<^sub>n.cases"}, and the list
   764   of mutual induction rules is called @{text
   765   "R\<^sub>1_\<dots>_R\<^sub>n.inducts"}.
   766 *}
   767 
   768 
   769 subsection {* Monotonicity theorems *}
   770 
   771 text {*
   772   Each theory contains a default set of theorems that are used in
   773   monotonicity proofs.  New rules can be added to this set via the
   774   @{attribute (HOL) mono} attribute.  The HOL theory @{text Inductive}
   775   shows how this is done.  In general, the following monotonicity
   776   theorems may be added:
   777 
   778   \begin{itemize}
   779 
   780   \item Theorems of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"}, for proving
   781   monotonicity of inductive definitions whose introduction rules have
   782   premises involving terms such as @{text "M R\<^sub>i t"}.
   783 
   784   \item Monotonicity theorems for logical operators, which are of the
   785   general form @{text "(\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> (\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> \<longrightarrow> \<dots>"}.  For example, in
   786   the case of the operator @{text "\<or>"}, the corresponding theorem is
   787   \[
   788   \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"}}
   789   \]
   790 
   791   \item De Morgan style equations for reasoning about the ``polarity''
   792   of expressions, e.g.
   793   \[
   794   @{prop "\<not> \<not> P \<longleftrightarrow> P"} \qquad\qquad
   795   @{prop "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"}
   796   \]
   797 
   798   \item Equations for reducing complex operators to more primitive
   799   ones whose monotonicity can easily be proved, e.g.
   800   \[
   801   @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> \<not> P \<or> Q"} \qquad\qquad
   802   @{prop "Ball A P \<equiv> \<forall>x. x \<in> A \<longrightarrow> P x"}
   803   \]
   804 
   805   \end{itemize}
   806 
   807   %FIXME: Example of an inductive definition
   808 *}
   809 
   810 
   811 section {* Arithmetic proof support *}
   812 
   813 text {*
   814   \begin{matharray}{rcl}
   815     @{method_def (HOL) arith} & : & \isarmeth \\
   816     @{method_def (HOL) arith_split} & : & \isaratt \\
   817   \end{matharray}
   818 
   819   The @{method (HOL) arith} method decides linear arithmetic problems
   820   (on types @{text nat}, @{text int}, @{text real}).  Any current
   821   facts are inserted into the goal before running the procedure.
   822 
   823   The @{method (HOL) arith_split} attribute declares case split rules
   824   to be expanded before the arithmetic procedure is invoked.
   825 
   826   Note that a simpler (but faster) version of arithmetic reasoning is
   827   already performed by the Simplifier.
   828 *}
   829 
   830 
   831 section {* Cases and induction: emulating tactic scripts \label{sec:hol-induct-tac} *}
   832 
   833 text {*
   834   The following important tactical tools of Isabelle/HOL have been
   835   ported to Isar.  These should be never used in proper proof texts!
   836 
   837   \begin{matharray}{rcl}
   838     @{method_def (HOL) case_tac}@{text "\<^sup>*"} & : & \isarmeth \\
   839     @{method_def (HOL) induct_tac}@{text "\<^sup>*"} & : & \isarmeth \\
   840     @{method_def (HOL) ind_cases}@{text "\<^sup>*"} & : & \isarmeth \\
   841     @{command_def (HOL) "inductive_cases"} & : & \isartrans{theory}{theory} \\
   842   \end{matharray}
   843 
   844   \begin{rail}
   845     'case\_tac' goalspec? term rule?
   846     ;
   847     'induct\_tac' goalspec? (insts * 'and') rule?
   848     ;
   849     'ind\_cases' (prop +) ('for' (name +)) ?
   850     ;
   851     'inductive\_cases' (thmdecl? (prop +) + 'and')
   852     ;
   853 
   854     rule: ('rule' ':' thmref)
   855     ;
   856   \end{rail}
   857 
   858   \begin{descr}
   859 
   860   \item [@{method (HOL) case_tac} and @{method (HOL) induct_tac}]
   861   admit to reason about inductive datatypes only (unless an
   862   alternative rule is given explicitly).  Furthermore, @{method (HOL)
   863   case_tac} does a classical case split on booleans; @{method (HOL)
   864   induct_tac} allows only variables to be given as instantiation.
   865   These tactic emulations feature both goal addressing and dynamic
   866   instantiation.  Note that named rule cases are \emph{not} provided
   867   as would be by the proper @{method induct} and @{method cases} proof
   868   methods (see \secref{sec:cases-induct}).
   869   
   870   \item [@{method (HOL) ind_cases} and @{command (HOL)
   871   "inductive_cases"}] provide an interface to the internal
   872   \texttt{mk_cases} operation.  Rules are simplified in an
   873   unrestricted forward manner.
   874 
   875   While @{method (HOL) ind_cases} is a proof method to apply the
   876   result immediately as elimination rules, @{command (HOL)
   877   "inductive_cases"} provides case split theorems at the theory level
   878   for later use.  The @{keyword "for"} argument of the @{method (HOL)
   879   ind_cases} method allows to specify a list of variables that should
   880   be generalized before applying the resulting rule.
   881 
   882   \end{descr}
   883 *}
   884 
   885 
   886 section {* Executable code *}
   887 
   888 text {*
   889   Isabelle/Pure provides two generic frameworks to support code
   890   generation from executable specifications.  Isabelle/HOL
   891   instantiates these mechanisms in a way that is amenable to end-user
   892   applications.
   893 
   894   One framework generates code from both functional and relational
   895   programs to SML.  See \cite{isabelle-HOL} for further information
   896   (this actually covers the new-style theory format as well).
   897 
   898   \begin{matharray}{rcl}
   899     @{command_def (HOL) "value"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
   900     @{command_def (HOL) "code_module"} & : & \isartrans{theory}{theory} \\
   901     @{command_def (HOL) "code_library"} & : & \isartrans{theory}{theory} \\
   902     @{command_def (HOL) "consts_code"} & : & \isartrans{theory}{theory} \\
   903     @{command_def (HOL) "types_code"} & : & \isartrans{theory}{theory} \\  
   904     @{attribute_def (HOL) code} & : & \isaratt \\
   905   \end{matharray}
   906 
   907   \begin{rail}
   908   'value' term
   909   ;
   910 
   911   ( 'code\_module' | 'code\_library' ) modespec ? name ? \\
   912     ( 'file' name ) ? ( 'imports' ( name + ) ) ? \\
   913     'contains' ( ( name '=' term ) + | term + )
   914   ;
   915 
   916   modespec: '(' ( name * ) ')'
   917   ;
   918 
   919   'consts\_code' (codespec +)
   920   ;
   921 
   922   codespec: const template attachment ?
   923   ;
   924 
   925   'types\_code' (tycodespec +)
   926   ;
   927 
   928   tycodespec: name template attachment ?
   929   ;
   930 
   931   const: term
   932   ;
   933 
   934   template: '(' string ')'
   935   ;
   936 
   937   attachment: 'attach' modespec ? verblbrace text verbrbrace
   938   ;
   939 
   940   'code' (name)?
   941   ;
   942   \end{rail}
   943 
   944   \begin{descr}
   945 
   946   \item [@{command (HOL) "value"}~@{text t}] evaluates and prints a
   947   term using the code generator.
   948 
   949   \end{descr}
   950 
   951   \medskip The other framework generates code from functional programs
   952   (including overloading using type classes) to SML \cite{SML}, OCaml
   953   \cite{OCaml} and Haskell \cite{haskell-revised-report}.
   954   Conceptually, code generation is split up in three steps:
   955   \emph{selection} of code theorems, \emph{translation} into an
   956   abstract executable view and \emph{serialization} to a specific
   957   \emph{target language}.  See \cite{isabelle-codegen} for an
   958   introduction on how to use it.
   959 
   960   \begin{matharray}{rcl}
   961     @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
   962     @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
   963     @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
   964     @{command_def (HOL) "code_datatype"} & : & \isartrans{theory}{theory} \\
   965     @{command_def (HOL) "code_const"} & : & \isartrans{theory}{theory} \\
   966     @{command_def (HOL) "code_type"} & : & \isartrans{theory}{theory} \\
   967     @{command_def (HOL) "code_class"} & : & \isartrans{theory}{theory} \\
   968     @{command_def (HOL) "code_instance"} & : & \isartrans{theory}{theory} \\
   969     @{command_def (HOL) "code_monad"} & : & \isartrans{theory}{theory} \\
   970     @{command_def (HOL) "code_reserved"} & : & \isartrans{theory}{theory} \\
   971     @{command_def (HOL) "code_include"} & : & \isartrans{theory}{theory} \\
   972     @{command_def (HOL) "code_modulename"} & : & \isartrans{theory}{theory} \\
   973     @{command_def (HOL) "code_exception"} & : & \isartrans{theory}{theory} \\
   974     @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
   975     @{attribute_def (HOL) code} & : & \isaratt \\
   976   \end{matharray}
   977 
   978   \begin{rail}
   979     'export\_code' ( constexpr + ) ? \\
   980       ( ( 'in' target ( 'module\_name' string ) ? \\
   981         ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
   982     ;
   983 
   984     'code\_thms' ( constexpr + ) ?
   985     ;
   986 
   987     'code\_deps' ( constexpr + ) ?
   988     ;
   989 
   990     const: term
   991     ;
   992 
   993     constexpr: ( const | 'name.*' | '*' )
   994     ;
   995 
   996     typeconstructor: nameref
   997     ;
   998 
   999     class: nameref
  1000     ;
  1001 
  1002     target: 'OCaml' | 'SML' | 'Haskell'
  1003     ;
  1004 
  1005     'code\_datatype' const +
  1006     ;
  1007 
  1008     'code\_const' (const + 'and') \\
  1009       ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
  1010     ;
  1011 
  1012     'code\_type' (typeconstructor + 'and') \\
  1013       ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
  1014     ;
  1015 
  1016     'code\_class' (class + 'and') \\
  1017       ( ( '(' target \\
  1018         ( ( string ('where' \\
  1019           ( const ( '==' | equiv ) string ) + ) ? ) ? + 'and' ) ')' ) + )
  1020     ;
  1021 
  1022     'code\_instance' (( typeconstructor '::' class ) + 'and') \\
  1023       ( ( '(' target ( '-' ? + 'and' ) ')' ) + )
  1024     ;
  1025 
  1026     'code\_monad' const const target
  1027     ;
  1028 
  1029     'code\_reserved' target ( string + )
  1030     ;
  1031 
  1032     'code\_include' target ( string ( string | '-') )
  1033     ;
  1034 
  1035     'code\_modulename' target ( ( string string ) + )
  1036     ;
  1037 
  1038     'code\_exception' ( const + )
  1039     ;
  1040 
  1041     syntax: string | ( 'infix' | 'infixl' | 'infixr' ) nat string
  1042     ;
  1043 
  1044     'code' ('func' | 'inline') ( 'del' )?
  1045     ;
  1046   \end{rail}
  1047 
  1048   \begin{descr}
  1049 
  1050   \item [@{command (HOL) "export_code"}] is the canonical interface
  1051   for generating and serializing code: for a given list of constants,
  1052   code is generated for the specified target languages.  Abstract code
  1053   is cached incrementally.  If no constant is given, the currently
  1054   cached code is serialized.  If no serialization instruction is
  1055   given, only abstract code is cached.
  1056 
  1057   Constants may be specified by giving them literally, referring to
  1058   all executable contants within a certain theory by giving @{text
  1059   "name.*"}, or referring to \emph{all} executable constants currently
  1060   available by giving @{text "*"}.
  1061 
  1062   By default, for each involved theory one corresponding name space
  1063   module is generated.  Alternativly, a module name may be specified
  1064   after the @{keyword "module_name"} keyword; then \emph{all} code is
  1065   placed in this module.
  1066 
  1067   For \emph{SML} and \emph{OCaml}, the file specification refers to a
  1068   single file; for \emph{Haskell}, it refers to a whole directory,
  1069   where code is generated in multiple files reflecting the module
  1070   hierarchy.  The file specification ``@{text "-"}'' denotes standard
  1071   output.  For \emph{SML}, omitting the file specification compiles
  1072   code internally in the context of the current ML session.
  1073 
  1074   Serializers take an optional list of arguments in parentheses.  For
  1075   \emph{Haskell} a module name prefix may be given using the ``@{text
  1076   "root:"}'' argument; ``@{text string_classes}'' adds a ``@{verbatim
  1077   "deriving (Read, Show)"}'' clause to each appropriate datatype
  1078   declaration.
  1079 
  1080   \item [@{command (HOL) "code_thms"}] prints a list of theorems
  1081   representing the corresponding program containing all given
  1082   constants; if no constants are given, the currently cached code
  1083   theorems are printed.
  1084 
  1085   \item [@{command (HOL) "code_deps"}] visualizes dependencies of
  1086   theorems representing the corresponding program containing all given
  1087   constants; if no constants are given, the currently cached code
  1088   theorems are visualized.
  1089 
  1090   \item [@{command (HOL) "code_datatype"}] specifies a constructor set
  1091   for a logical type.
  1092 
  1093   \item [@{command (HOL) "code_const"}] associates a list of constants
  1094   with target-specific serializations; omitting a serialization
  1095   deletes an existing serialization.
  1096 
  1097   \item [@{command (HOL) "code_type"}] associates a list of type
  1098   constructors with target-specific serializations; omitting a
  1099   serialization deletes an existing serialization.
  1100 
  1101   \item [@{command (HOL) "code_class"}] associates a list of classes
  1102   with target-specific class names; in addition, constants associated
  1103   with this class may be given target-specific names used for instance
  1104   declarations; omitting a serialization deletes an existing
  1105   serialization.  This applies only to \emph{Haskell}.
  1106 
  1107   \item [@{command (HOL) "code_instance"}] declares a list of type
  1108   constructor / class instance relations as ``already present'' for a
  1109   given target.  Omitting a ``@{text "-"}'' deletes an existing
  1110   ``already present'' declaration.  This applies only to
  1111   \emph{Haskell}.
  1112 
  1113   \item [@{command (HOL) "code_monad"}] provides an auxiliary
  1114   mechanism to generate monadic code.
  1115 
  1116   \item [@{command (HOL) "code_reserved"}] declares a list of names as
  1117   reserved for a given target, preventing it to be shadowed by any
  1118   generated code.
  1119 
  1120   \item [@{command (HOL) "code_include"}] adds arbitrary named content
  1121   (``include'') to generated code.  A as last argument ``@{text "-"}''
  1122   will remove an already added ``include''.
  1123 
  1124   \item [@{command (HOL) "code_modulename"}] declares aliasings from
  1125   one module name onto another.
  1126 
  1127   \item [@{command (HOL) "code_exception"}] declares constants which
  1128   are not required to have a definition by a defining equations; these
  1129   are mapped on exceptions instead.
  1130 
  1131   \item [@{attribute (HOL) code}~@{text func}] explicitly selects (or
  1132   with option ``@{text "del:"}'' deselects) a defining equation for
  1133   code generation.  Usually packages introducing defining equations
  1134   provide a resonable default setup for selection.
  1135 
  1136   \item [@{attribute (HOL) code}@{text inline}] declares (or with
  1137   option ``@{text "del:"}'' removes) inlining theorems which are
  1138   applied as rewrite rules to any defining equation during
  1139   preprocessing.
  1140 
  1141   \item [@{command (HOL) "print_codesetup"}] gives an overview on
  1142   selected defining equations, code generator datatypes and
  1143   preprocessor setup.
  1144 
  1145   \end{descr}
  1146 *}
  1147 
  1148 end
  1149