doc-src/IsarRef/Thy/HOL_Specific.thy
author wenzelm
Fri, 03 Apr 2009 18:03:29 +0200
changeset 30865 5106e13d400f
parent 30863 5dc392a59bb7
child 31248 03a35fbc9dc6
permissions -rw-r--r--
fixed formal markup;
     1 theory HOL_Specific
     2 imports Main
     3 begin
     4 
     5 chapter {* Isabelle/HOL \label{ch:hol} *}
     6 
     7 section {* Primitive types \label{sec:hol-typedef} *}
     8 
     9 text {*
    10   \begin{matharray}{rcl}
    11     @{command_def (HOL) "typedecl"} & : & @{text "theory \<rightarrow> theory"} \\
    12     @{command_def (HOL) "typedef"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
    13   \end{matharray}
    14 
    15   \begin{rail}
    16     'typedecl' typespec infix?
    17     ;
    18     'typedef' altname? abstype '=' repset
    19     ;
    20 
    21     altname: '(' (name | 'open' | 'open' name) ')'
    22     ;
    23     abstype: typespec infix?
    24     ;
    25     repset: term ('morphisms' name name)?
    26     ;
    27   \end{rail}
    28 
    29   \begin{description}
    30   
    31   \item @{command (HOL) "typedecl"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"} is similar
    32   to the original @{command "typedecl"} of Isabelle/Pure (see
    33   \secref{sec:types-pure}), but also declares type arity @{text "t ::
    34   (type, \<dots>, type) type"}, making @{text t} an actual HOL type
    35   constructor.  %FIXME check, update
    36   
    37   \item @{command (HOL) "typedef"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = A"} sets up
    38   a goal stating non-emptiness of the set @{text A}.  After finishing
    39   the proof, the theory will be augmented by a Gordon/HOL-style type
    40   definition, which establishes a bijection between the representing
    41   set @{text A} and the new type @{text t}.
    42   
    43   Technically, @{command (HOL) "typedef"} defines both a type @{text
    44   t} and a set (term constant) of the same name (an alternative base
    45   name may be given in parentheses).  The injection from type to set
    46   is called @{text Rep_t}, its inverse @{text Abs_t} (this may be
    47   changed via an explicit @{keyword (HOL) "morphisms"} declaration).
    48   
    49   Theorems @{text Rep_t}, @{text Rep_t_inverse}, and @{text
    50   Abs_t_inverse} provide the most basic characterization as a
    51   corresponding injection/surjection pair (in both directions).  Rules
    52   @{text Rep_t_inject} and @{text Abs_t_inject} provide a slightly
    53   more convenient view on the injectivity part, suitable for automated
    54   proof tools (e.g.\ in @{attribute simp} or @{attribute iff}
    55   declarations).  Rules @{text Rep_t_cases}/@{text Rep_t_induct}, and
    56   @{text Abs_t_cases}/@{text Abs_t_induct} provide alternative views
    57   on surjectivity; these are already declared as set or type rules for
    58   the generic @{method cases} and @{method induct} methods.
    59   
    60   An alternative name may be specified in parentheses; the default is
    61   to use @{text t} as indicated before.  The ``@{text "(open)"}''
    62   declaration suppresses a separate constant definition for the
    63   representing set.
    64 
    65   \end{description}
    66 
    67   Note that raw type declarations are rarely used in practice; the
    68   main application is with experimental (or even axiomatic!) theory
    69   fragments.  Instead of primitive HOL type definitions, user-level
    70   theories usually refer to higher-level packages such as @{command
    71   (HOL) "record"} (see \secref{sec:hol-record}) or @{command (HOL)
    72   "datatype"} (see \secref{sec:hol-datatype}).
    73 *}
    74 
    75 
    76 section {* Adhoc tuples *}
    77 
    78 text {*
    79   \begin{matharray}{rcl}
    80     @{attribute (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\
    81   \end{matharray}
    82 
    83   \begin{rail}
    84     'split\_format' (((name *) + 'and') | ('(' 'complete' ')'))
    85     ;
    86   \end{rail}
    87 
    88   \begin{description}
    89   
    90   \item @{attribute (HOL) split_format}~@{text "p\<^sub>1 \<dots> p\<^sub>m \<AND> \<dots>
    91   \<AND> q\<^sub>1 \<dots> q\<^sub>n"} puts expressions of low-level tuple types into
    92   canonical form as specified by the arguments given; the @{text i}-th
    93   collection of arguments refers to occurrences in premise @{text i}
    94   of the rule.  The ``@{text "(complete)"}'' option causes \emph{all}
    95   arguments in function applications to be represented canonically
    96   according to their tuple type structure.
    97 
    98   Note that these operations tend to invent funny names for new local
    99   parameters to be introduced.
   100 
   101   \end{description}
   102 *}
   103 
   104 
   105 section {* Records \label{sec:hol-record} *}
   106 
   107 text {*
   108   In principle, records merely generalize the concept of tuples, where
   109   components may be addressed by labels instead of just position.  The
   110   logical infrastructure of records in Isabelle/HOL is slightly more
   111   advanced, though, supporting truly extensible record schemes.  This
   112   admits operations that are polymorphic with respect to record
   113   extension, yielding ``object-oriented'' effects like (single)
   114   inheritance.  See also \cite{NaraschewskiW-TPHOLs98} for more
   115   details on object-oriented verification and record subtyping in HOL.
   116 *}
   117 
   118 
   119 subsection {* Basic concepts *}
   120 
   121 text {*
   122   Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records
   123   at the level of terms and types.  The notation is as follows:
   124 
   125   \begin{center}
   126   \begin{tabular}{l|l|l}
   127     & record terms & record types \\ \hline
   128     fixed & @{text "\<lparr>x = a, y = b\<rparr>"} & @{text "\<lparr>x :: A, y :: B\<rparr>"} \\
   129     schematic & @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} &
   130       @{text "\<lparr>x :: A, y :: B, \<dots> :: M\<rparr>"} \\
   131   \end{tabular}
   132   \end{center}
   133 
   134   \noindent The ASCII representation of @{text "\<lparr>x = a\<rparr>"} is @{text
   135   "(| x = a |)"}.
   136 
   137   A fixed record @{text "\<lparr>x = a, y = b\<rparr>"} has field @{text x} of value
   138   @{text a} and field @{text y} of value @{text b}.  The corresponding
   139   type is @{text "\<lparr>x :: A, y :: B\<rparr>"}, assuming that @{text "a :: A"}
   140   and @{text "b :: B"}.
   141 
   142   A record scheme like @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} contains fields
   143   @{text x} and @{text y} as before, but also possibly further fields
   144   as indicated by the ``@{text "\<dots>"}'' notation (which is actually part
   145   of the syntax).  The improper field ``@{text "\<dots>"}'' of a record
   146   scheme is called the \emph{more part}.  Logically it is just a free
   147   variable, which is occasionally referred to as ``row variable'' in
   148   the literature.  The more part of a record scheme may be
   149   instantiated by zero or more further components.  For example, the
   150   previous scheme may get instantiated to @{text "\<lparr>x = a, y = b, z =
   151   c, \<dots> = m'\<rparr>"}, where @{text m'} refers to a different more part.
   152   Fixed records are special instances of record schemes, where
   153   ``@{text "\<dots>"}'' is properly terminated by the @{text "() :: unit"}
   154   element.  In fact, @{text "\<lparr>x = a, y = b\<rparr>"} is just an abbreviation
   155   for @{text "\<lparr>x = a, y = b, \<dots> = ()\<rparr>"}.
   156   
   157   \medskip Two key observations make extensible records in a simply
   158   typed language like HOL work out:
   159 
   160   \begin{enumerate}
   161 
   162   \item the more part is internalized, as a free term or type
   163   variable,
   164 
   165   \item field names are externalized, they cannot be accessed within
   166   the logic as first-class values.
   167 
   168   \end{enumerate}
   169 
   170   \medskip In Isabelle/HOL record types have to be defined explicitly,
   171   fixing their field names and types, and their (optional) parent
   172   record.  Afterwards, records may be formed using above syntax, while
   173   obeying the canonical order of fields as given by their declaration.
   174   The record package provides several standard operations like
   175   selectors and updates.  The common setup for various generic proof
   176   tools enable succinct reasoning patterns.  See also the Isabelle/HOL
   177   tutorial \cite{isabelle-hol-book} for further instructions on using
   178   records in practice.
   179 *}
   180 
   181 
   182 subsection {* Record specifications *}
   183 
   184 text {*
   185   \begin{matharray}{rcl}
   186     @{command_def (HOL) "record"} & : & @{text "theory \<rightarrow> theory"} \\
   187   \end{matharray}
   188 
   189   \begin{rail}
   190     'record' typespec '=' (type '+')? (constdecl +)
   191     ;
   192   \end{rail}
   193 
   194   \begin{description}
   195 
   196   \item @{command (HOL) "record"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t = \<tau> + c\<^sub>1 :: \<sigma>\<^sub>1
   197   \<dots> c\<^sub>n :: \<sigma>\<^sub>n"} defines extensible record type @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"},
   198   derived from the optional parent record @{text "\<tau>"} by adding new
   199   field components @{text "c\<^sub>i :: \<sigma>\<^sub>i"} etc.
   200 
   201   The type variables of @{text "\<tau>"} and @{text "\<sigma>\<^sub>i"} need to be
   202   covered by the (distinct) parameters @{text "\<alpha>\<^sub>1, \<dots>,
   203   \<alpha>\<^sub>m"}.  Type constructor @{text t} has to be new, while @{text
   204   \<tau>} needs to specify an instance of an existing record type.  At
   205   least one new field @{text "c\<^sub>i"} has to be specified.
   206   Basically, field names need to belong to a unique record.  This is
   207   not a real restriction in practice, since fields are qualified by
   208   the record name internally.
   209 
   210   The parent record specification @{text \<tau>} is optional; if omitted
   211   @{text t} becomes a root record.  The hierarchy of all records
   212   declared within a theory context forms a forest structure, i.e.\ a
   213   set of trees starting with a root record each.  There is no way to
   214   merge multiple parent records!
   215 
   216   For convenience, @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} is made a
   217   type abbreviation for the fixed record type @{text "\<lparr>c\<^sub>1 ::
   218   \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n\<rparr>"}, likewise is @{text
   219   "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m, \<zeta>) t_scheme"} made an abbreviation for
   220   @{text "\<lparr>c\<^sub>1 :: \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n, \<dots> ::
   221   \<zeta>\<rparr>"}.
   222 
   223   \end{description}
   224 *}
   225 
   226 
   227 subsection {* Record operations *}
   228 
   229 text {*
   230   Any record definition of the form presented above produces certain
   231   standard operations.  Selectors and updates are provided for any
   232   field, including the improper one ``@{text more}''.  There are also
   233   cumulative record constructor functions.  To simplify the
   234   presentation below, we assume for now that @{text "(\<alpha>\<^sub>1, \<dots>,
   235   \<alpha>\<^sub>m) t"} is a root record with fields @{text "c\<^sub>1 ::
   236   \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n"}.
   237 
   238   \medskip \textbf{Selectors} and \textbf{updates} are available for
   239   any field (including ``@{text more}''):
   240 
   241   \begin{matharray}{lll}
   242     @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
   243     @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
   244   \end{matharray}
   245 
   246   There is special syntax for application of updates: @{text "r\<lparr>x :=
   247   a\<rparr>"} abbreviates term @{text "x_update a r"}.  Further notation for
   248   repeated updates is also available: @{text "r\<lparr>x := a\<rparr>\<lparr>y := b\<rparr>\<lparr>z :=
   249   c\<rparr>"} may be written @{text "r\<lparr>x := a, y := b, z := c\<rparr>"}.  Note that
   250   because of postfix notation the order of fields shown here is
   251   reverse than in the actual term.  Since repeated updates are just
   252   function applications, fields may be freely permuted in @{text "\<lparr>x
   253   := a, y := b, z := c\<rparr>"}, as far as logical equality is concerned.
   254   Thus commutativity of independent updates can be proven within the
   255   logic for any two fields, but not as a general theorem.
   256 
   257   \medskip The \textbf{make} operation provides a cumulative record
   258   constructor function:
   259 
   260   \begin{matharray}{lll}
   261     @{text "t.make"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   262   \end{matharray}
   263 
   264   \medskip We now reconsider the case of non-root records, which are
   265   derived of some parent.  In general, the latter may depend on
   266   another parent as well, resulting in a list of \emph{ancestor
   267   records}.  Appending the lists of fields of all ancestors results in
   268   a certain field prefix.  The record package automatically takes care
   269   of this by lifting operations over this context of ancestor fields.
   270   Assuming that @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} has ancestor
   271   fields @{text "b\<^sub>1 :: \<rho>\<^sub>1, \<dots>, b\<^sub>k :: \<rho>\<^sub>k"},
   272   the above record operations will get the following types:
   273 
   274   \medskip
   275   \begin{tabular}{lll}
   276     @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
   277     @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow> 
   278       \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow>
   279       \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
   280     @{text "t.make"} & @{text "::"} & @{text "\<rho>\<^sub>1 \<Rightarrow> \<dots> \<rho>\<^sub>k \<Rightarrow> \<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow>
   281       \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   282   \end{tabular}
   283   \medskip
   284 
   285   \noindent Some further operations address the extension aspect of a
   286   derived record scheme specifically: @{text "t.fields"} produces a
   287   record fragment consisting of exactly the new fields introduced here
   288   (the result may serve as a more part elsewhere); @{text "t.extend"}
   289   takes a fixed record and adds a given more part; @{text
   290   "t.truncate"} restricts a record scheme to a fixed record.
   291 
   292   \medskip
   293   \begin{tabular}{lll}
   294     @{text "t.fields"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   295     @{text "t.extend"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr> \<Rightarrow>
   296       \<zeta> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
   297     @{text "t.truncate"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   298   \end{tabular}
   299   \medskip
   300 
   301   \noindent Note that @{text "t.make"} and @{text "t.fields"} coincide
   302   for root records.
   303 *}
   304 
   305 
   306 subsection {* Derived rules and proof tools *}
   307 
   308 text {*
   309   The record package proves several results internally, declaring
   310   these facts to appropriate proof tools.  This enables users to
   311   reason about record structures quite conveniently.  Assume that
   312   @{text t} is a record type as specified above.
   313 
   314   \begin{enumerate}
   315   
   316   \item Standard conversions for selectors or updates applied to
   317   record constructor terms are made part of the default Simplifier
   318   context; thus proofs by reduction of basic operations merely require
   319   the @{method simp} method without further arguments.  These rules
   320   are available as @{text "t.simps"}, too.
   321   
   322   \item Selectors applied to updated records are automatically reduced
   323   by an internal simplification procedure, which is also part of the
   324   standard Simplifier setup.
   325 
   326   \item Inject equations of a form analogous to @{prop "(x, y) = (x',
   327   y') \<equiv> x = x' \<and> y = y'"} are declared to the Simplifier and Classical
   328   Reasoner as @{attribute iff} rules.  These rules are available as
   329   @{text "t.iffs"}.
   330 
   331   \item The introduction rule for record equality analogous to @{text
   332   "x r = x r' \<Longrightarrow> y r = y r' \<dots> \<Longrightarrow> r = r'"} is declared to the Simplifier,
   333   and as the basic rule context as ``@{attribute intro}@{text "?"}''.
   334   The rule is called @{text "t.equality"}.
   335 
   336   \item Representations of arbitrary record expressions as canonical
   337   constructor terms are provided both in @{method cases} and @{method
   338   induct} format (cf.\ the generic proof methods of the same name,
   339   \secref{sec:cases-induct}).  Several variations are available, for
   340   fixed records, record schemes, more parts etc.
   341   
   342   The generic proof methods are sufficiently smart to pick the most
   343   sensible rule according to the type of the indicated record
   344   expression: users just need to apply something like ``@{text "(cases
   345   r)"}'' to a certain proof problem.
   346 
   347   \item The derived record operations @{text "t.make"}, @{text
   348   "t.fields"}, @{text "t.extend"}, @{text "t.truncate"} are \emph{not}
   349   treated automatically, but usually need to be expanded by hand,
   350   using the collective fact @{text "t.defs"}.
   351 
   352   \end{enumerate}
   353 *}
   354 
   355 
   356 section {* Datatypes \label{sec:hol-datatype} *}
   357 
   358 text {*
   359   \begin{matharray}{rcl}
   360     @{command_def (HOL) "datatype"} & : & @{text "theory \<rightarrow> theory"} \\
   361   @{command_def (HOL) "rep_datatype"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   362   \end{matharray}
   363 
   364   \begin{rail}
   365     'datatype' (dtspec + 'and')
   366     ;
   367     'rep\_datatype' ('(' (name +) ')')? (term +)
   368     ;
   369 
   370     dtspec: parname? typespec infix? '=' (cons + '|')
   371     ;
   372     cons: name (type *) mixfix?
   373   \end{rail}
   374 
   375   \begin{description}
   376 
   377   \item @{command (HOL) "datatype"} defines inductive datatypes in
   378   HOL.
   379 
   380   \item @{command (HOL) "rep_datatype"} represents existing types as
   381   inductive ones, generating the standard infrastructure of derived
   382   concepts (primitive recursion etc.).
   383 
   384   \end{description}
   385 
   386   The induction and exhaustion theorems generated provide case names
   387   according to the constructors involved, while parameters are named
   388   after the types (see also \secref{sec:cases-induct}).
   389 
   390   See \cite{isabelle-HOL} for more details on datatypes, but beware of
   391   the old-style theory syntax being used there!  Apart from proper
   392   proof methods for case-analysis and induction, there are also
   393   emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL)
   394   induct_tac} available, see \secref{sec:hol-induct-tac}; these admit
   395   to refer directly to the internal structure of subgoals (including
   396   internally bound parameters).
   397 *}
   398 
   399 
   400 section {* Recursive functions \label{sec:recursion} *}
   401 
   402 text {*
   403   \begin{matharray}{rcl}
   404     @{command_def (HOL) "primrec"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   405     @{command_def (HOL) "fun"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   406     @{command_def (HOL) "function"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
   407     @{command_def (HOL) "termination"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
   408   \end{matharray}
   409 
   410   \begin{rail}
   411     'primrec' target? fixes 'where' equations
   412     ;
   413     equations: (thmdecl? prop + '|')
   414     ;
   415     ('fun' | 'function') target? functionopts? fixes 'where' clauses
   416     ;
   417     clauses: (thmdecl? prop ('(' 'otherwise' ')')? + '|')
   418     ;
   419     functionopts: '(' (('sequential' | 'domintros' | 'tailrec' | 'default' term) + ',') ')'
   420     ;
   421     'termination' ( term )?
   422   \end{rail}
   423 
   424   \begin{description}
   425 
   426   \item @{command (HOL) "primrec"} defines primitive recursive
   427   functions over datatypes, see also \cite{isabelle-HOL}.
   428 
   429   \item @{command (HOL) "function"} defines functions by general
   430   wellfounded recursion. A detailed description with examples can be
   431   found in \cite{isabelle-function}. The function is specified by a
   432   set of (possibly conditional) recursive equations with arbitrary
   433   pattern matching. The command generates proof obligations for the
   434   completeness and the compatibility of patterns.
   435 
   436   The defined function is considered partial, and the resulting
   437   simplification rules (named @{text "f.psimps"}) and induction rule
   438   (named @{text "f.pinduct"}) are guarded by a generated domain
   439   predicate @{text "f_dom"}. The @{command (HOL) "termination"}
   440   command can then be used to establish that the function is total.
   441 
   442   \item @{command (HOL) "fun"} is a shorthand notation for ``@{command
   443   (HOL) "function"}~@{text "(sequential)"}, followed by automated
   444   proof attempts regarding pattern matching and termination.  See
   445   \cite{isabelle-function} for further details.
   446 
   447   \item @{command (HOL) "termination"}~@{text f} commences a
   448   termination proof for the previously defined function @{text f}.  If
   449   this is omitted, the command refers to the most recent function
   450   definition.  After the proof is closed, the recursive equations and
   451   the induction principle is established.
   452 
   453   \end{description}
   454 
   455   %FIXME check
   456 
   457   Recursive definitions introduced by the @{command (HOL) "function"}
   458   command accommodate
   459   reasoning by induction (cf.\ \secref{sec:cases-induct}): rule @{text
   460   "c.induct"} (where @{text c} is the name of the function definition)
   461   refers to a specific induction rule, with parameters named according
   462   to the user-specified equations.
   463   For the @{command (HOL) "primrec"} the induction principle coincides
   464   with structural recursion on the datatype the recursion is carried
   465   out.
   466   Case names of @{command (HOL)
   467   "primrec"} are that of the datatypes involved, while those of
   468   @{command (HOL) "function"} are numbered (starting from 1).
   469 
   470   The equations provided by these packages may be referred later as
   471   theorem list @{text "f.simps"}, where @{text f} is the (collective)
   472   name of the functions defined.  Individual equations may be named
   473   explicitly as well.
   474 
   475   The @{command (HOL) "function"} command accepts the following
   476   options.
   477 
   478   \begin{description}
   479 
   480   \item @{text sequential} enables a preprocessor which disambiguates
   481   overlapping patterns by making them mutually disjoint.  Earlier
   482   equations take precedence over later ones.  This allows to give the
   483   specification in a format very similar to functional programming.
   484   Note that the resulting simplification and induction rules
   485   correspond to the transformed specification, not the one given
   486   originally. This usually means that each equation given by the user
   487   may result in several theroems.  Also note that this automatic
   488   transformation only works for ML-style datatype patterns.
   489 
   490   \item @{text domintros} enables the automated generation of
   491   introduction rules for the domain predicate. While mostly not
   492   needed, they can be helpful in some proofs about partial functions.
   493 
   494   \item @{text tailrec} generates the unconstrained recursive
   495   equations even without a termination proof, provided that the
   496   function is tail-recursive. This currently only works
   497 
   498   \item @{text "default d"} allows to specify a default value for a
   499   (partial) function, which will ensure that @{text "f x = d x"}
   500   whenever @{text "x \<notin> f_dom"}.
   501 
   502   \end{description}
   503 *}
   504 
   505 
   506 subsection {* Proof methods related to recursive definitions *}
   507 
   508 text {*
   509   \begin{matharray}{rcl}
   510     @{method_def (HOL) pat_completeness} & : & @{text method} \\
   511     @{method_def (HOL) relation} & : & @{text method} \\
   512     @{method_def (HOL) lexicographic_order} & : & @{text method} \\
   513   \end{matharray}
   514 
   515   \begin{rail}
   516     'relation' term
   517     ;
   518     'lexicographic\_order' (clasimpmod *)
   519     ;
   520   \end{rail}
   521 
   522   \begin{description}
   523 
   524   \item @{method (HOL) pat_completeness} is a specialized method to
   525   solve goals regarding the completeness of pattern matching, as
   526   required by the @{command (HOL) "function"} package (cf.\
   527   \cite{isabelle-function}).
   528 
   529   \item @{method (HOL) relation}~@{text R} introduces a termination
   530   proof using the relation @{text R}.  The resulting proof state will
   531   contain goals expressing that @{text R} is wellfounded, and that the
   532   arguments of recursive calls decrease with respect to @{text R}.
   533   Usually, this method is used as the initial proof step of manual
   534   termination proofs.
   535 
   536   \item @{method (HOL) "lexicographic_order"} attempts a fully
   537   automated termination proof by searching for a lexicographic
   538   combination of size measures on the arguments of the function. The
   539   method accepts the same arguments as the @{method auto} method,
   540   which it uses internally to prove local descents.  The same context
   541   modifiers as for @{method auto} are accepted, see
   542   \secref{sec:clasimp}.
   543 
   544   In case of failure, extensive information is printed, which can help
   545   to analyse the situation (cf.\ \cite{isabelle-function}).
   546 
   547   \end{description}
   548 *}
   549 
   550 
   551 subsection {* Old-style recursive function definitions (TFL) *}
   552 
   553 text {*
   554   The old TFL commands @{command (HOL) "recdef"} and @{command (HOL)
   555   "recdef_tc"} for defining recursive are mostly obsolete; @{command
   556   (HOL) "function"} or @{command (HOL) "fun"} should be used instead.
   557 
   558   \begin{matharray}{rcl}
   559     @{command_def (HOL) "recdef"} & : & @{text "theory \<rightarrow> theory)"} \\
   560     @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   561   \end{matharray}
   562 
   563   \begin{rail}
   564     'recdef' ('(' 'permissive' ')')? \\ name term (prop +) hints?
   565     ;
   566     recdeftc thmdecl? tc
   567     ;
   568     hints: '(' 'hints' (recdefmod *) ')'
   569     ;
   570     recdefmod: (('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del') ':' thmrefs) | clasimpmod
   571     ;
   572     tc: nameref ('(' nat ')')?
   573     ;
   574   \end{rail}
   575 
   576   \begin{description}
   577   
   578   \item @{command (HOL) "recdef"} defines general well-founded
   579   recursive functions (using the TFL package), see also
   580   \cite{isabelle-HOL}.  The ``@{text "(permissive)"}'' option tells
   581   TFL to recover from failed proof attempts, returning unfinished
   582   results.  The @{text recdef_simp}, @{text recdef_cong}, and @{text
   583   recdef_wf} hints refer to auxiliary rules to be used in the internal
   584   automated proof process of TFL.  Additional @{syntax clasimpmod}
   585   declarations (cf.\ \secref{sec:clasimp}) may be given to tune the
   586   context of the Simplifier (cf.\ \secref{sec:simplifier}) and
   587   Classical reasoner (cf.\ \secref{sec:classical}).
   588   
   589   \item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the
   590   proof for leftover termination condition number @{text i} (default
   591   1) as generated by a @{command (HOL) "recdef"} definition of
   592   constant @{text c}.
   593   
   594   Note that in most cases, @{command (HOL) "recdef"} is able to finish
   595   its internal proofs without manual intervention.
   596 
   597   \end{description}
   598 
   599   \medskip Hints for @{command (HOL) "recdef"} may be also declared
   600   globally, using the following attributes.
   601 
   602   \begin{matharray}{rcl}
   603     @{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\
   604     @{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\
   605     @{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\
   606   \end{matharray}
   607 
   608   \begin{rail}
   609     ('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del')
   610     ;
   611   \end{rail}
   612 *}
   613 
   614 
   615 section {* Inductive and coinductive definitions \label{sec:hol-inductive} *}
   616 
   617 text {*
   618   An \textbf{inductive definition} specifies the least predicate (or
   619   set) @{text R} closed under given rules: applying a rule to elements
   620   of @{text R} yields a result within @{text R}.  For example, a
   621   structural operational semantics is an inductive definition of an
   622   evaluation relation.
   623 
   624   Dually, a \textbf{coinductive definition} specifies the greatest
   625   predicate~/ set @{text R} that is consistent with given rules: every
   626   element of @{text R} can be seen as arising by applying a rule to
   627   elements of @{text R}.  An important example is using bisimulation
   628   relations to formalise equivalence of processes and infinite data
   629   structures.
   630 
   631   \medskip The HOL package is related to the ZF one, which is
   632   described in a separate paper,\footnote{It appeared in CADE
   633   \cite{paulson-CADE}; a longer version is distributed with Isabelle.}
   634   which you should refer to in case of difficulties.  The package is
   635   simpler than that of ZF thanks to implicit type-checking in HOL.
   636   The types of the (co)inductive predicates (or sets) determine the
   637   domain of the fixedpoint definition, and the package does not have
   638   to use inference rules for type-checking.
   639 
   640   \begin{matharray}{rcl}
   641     @{command_def (HOL) "inductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   642     @{command_def (HOL) "inductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   643     @{command_def (HOL) "coinductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   644     @{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   645     @{attribute_def (HOL) mono} & : & @{text attribute} \\
   646   \end{matharray}
   647 
   648   \begin{rail}
   649     ('inductive' | 'inductive\_set' | 'coinductive' | 'coinductive\_set') target? fixes ('for' fixes)? \\
   650     ('where' clauses)? ('monos' thmrefs)?
   651     ;
   652     clauses: (thmdecl? prop + '|')
   653     ;
   654     'mono' (() | 'add' | 'del')
   655     ;
   656   \end{rail}
   657 
   658   \begin{description}
   659 
   660   \item @{command (HOL) "inductive"} and @{command (HOL)
   661   "coinductive"} define (co)inductive predicates from the
   662   introduction rules given in the @{keyword "where"} part.  The
   663   optional @{keyword "for"} part contains a list of parameters of the
   664   (co)inductive predicates that remain fixed throughout the
   665   definition.  The optional @{keyword "monos"} section contains
   666   \emph{monotonicity theorems}, which are required for each operator
   667   applied to a recursive set in the introduction rules.  There
   668   \emph{must} be a theorem of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"},
   669   for each premise @{text "M R\<^sub>i t"} in an introduction rule!
   670 
   671   \item @{command (HOL) "inductive_set"} and @{command (HOL)
   672   "coinductive_set"} are wrappers for to the previous commands,
   673   allowing the definition of (co)inductive sets.
   674 
   675   \item @{attribute (HOL) mono} declares monotonicity rules.  These
   676   rule are involved in the automated monotonicity proof of @{command
   677   (HOL) "inductive"}.
   678 
   679   \end{description}
   680 *}
   681 
   682 
   683 subsection {* Derived rules *}
   684 
   685 text {*
   686   Each (co)inductive definition @{text R} adds definitions to the
   687   theory and also proves some theorems:
   688 
   689   \begin{description}
   690 
   691   \item @{text R.intros} is the list of introduction rules as proven
   692   theorems, for the recursive predicates (or sets).  The rules are
   693   also available individually, using the names given them in the
   694   theory file;
   695 
   696   \item @{text R.cases} is the case analysis (or elimination) rule;
   697 
   698   \item @{text R.induct} or @{text R.coinduct} is the (co)induction
   699   rule.
   700 
   701   \end{description}
   702 
   703   When several predicates @{text "R\<^sub>1, \<dots>, R\<^sub>n"} are
   704   defined simultaneously, the list of introduction rules is called
   705   @{text "R\<^sub>1_\<dots>_R\<^sub>n.intros"}, the case analysis rules are
   706   called @{text "R\<^sub>1.cases, \<dots>, R\<^sub>n.cases"}, and the list
   707   of mutual induction rules is called @{text
   708   "R\<^sub>1_\<dots>_R\<^sub>n.inducts"}.
   709 *}
   710 
   711 
   712 subsection {* Monotonicity theorems *}
   713 
   714 text {*
   715   Each theory contains a default set of theorems that are used in
   716   monotonicity proofs.  New rules can be added to this set via the
   717   @{attribute (HOL) mono} attribute.  The HOL theory @{text Inductive}
   718   shows how this is done.  In general, the following monotonicity
   719   theorems may be added:
   720 
   721   \begin{itemize}
   722 
   723   \item Theorems of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"}, for proving
   724   monotonicity of inductive definitions whose introduction rules have
   725   premises involving terms such as @{text "M R\<^sub>i t"}.
   726 
   727   \item Monotonicity theorems for logical operators, which are of the
   728   general form @{text "(\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> (\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> \<longrightarrow> \<dots>"}.  For example, in
   729   the case of the operator @{text "\<or>"}, the corresponding theorem is
   730   \[
   731   \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"}}
   732   \]
   733 
   734   \item De Morgan style equations for reasoning about the ``polarity''
   735   of expressions, e.g.
   736   \[
   737   @{prop "\<not> \<not> P \<longleftrightarrow> P"} \qquad\qquad
   738   @{prop "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"}
   739   \]
   740 
   741   \item Equations for reducing complex operators to more primitive
   742   ones whose monotonicity can easily be proved, e.g.
   743   \[
   744   @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> \<not> P \<or> Q"} \qquad\qquad
   745   @{prop "Ball A P \<equiv> \<forall>x. x \<in> A \<longrightarrow> P x"}
   746   \]
   747 
   748   \end{itemize}
   749 
   750   %FIXME: Example of an inductive definition
   751 *}
   752 
   753 
   754 section {* Arithmetic proof support *}
   755 
   756 text {*
   757   \begin{matharray}{rcl}
   758     @{method_def (HOL) arith} & : & @{text method} \\
   759     @{attribute_def (HOL) arith} & : & @{text attribute} \\
   760     @{attribute_def (HOL) arith_split} & : & @{text attribute} \\
   761   \end{matharray}
   762 
   763   The @{method (HOL) arith} method decides linear arithmetic problems
   764   (on types @{text nat}, @{text int}, @{text real}).  Any current
   765   facts are inserted into the goal before running the procedure.
   766 
   767   The @{attribute (HOL) arith} attribute declares facts that are
   768   always supplied to the arithmetic provers implicitly.
   769 
   770   The @{attribute (HOL) arith_split} attribute declares case split
   771   rules to be expanded before @{method (HOL) arith} is invoked.
   772 
   773   Note that a simpler (but faster) arithmetic prover is
   774   already invoked by the Simplifier.
   775 *}
   776 
   777 
   778 section {* Intuitionistic proof search *}
   779 
   780 text {*
   781   \begin{matharray}{rcl}
   782     @{method_def (HOL) iprover} & : & @{text method} \\
   783   \end{matharray}
   784 
   785   \begin{rail}
   786     'iprover' ('!' ?) (rulemod *)
   787     ;
   788   \end{rail}
   789 
   790   The @{method (HOL) iprover} method performs intuitionistic proof
   791   search, depending on specifically declared rules from the context,
   792   or given as explicit arguments.  Chained facts are inserted into the
   793   goal before commencing proof search; ``@{method (HOL) iprover}@{text
   794   "!"}''  means to include the current @{fact prems} as well.
   795   
   796   Rules need to be classified as @{attribute (Pure) intro},
   797   @{attribute (Pure) elim}, or @{attribute (Pure) dest}; here the
   798   ``@{text "!"}'' indicator refers to ``safe'' rules, which may be
   799   applied aggressively (without considering back-tracking later).
   800   Rules declared with ``@{text "?"}'' are ignored in proof search (the
   801   single-step @{method rule} method still observes these).  An
   802   explicit weight annotation may be given as well; otherwise the
   803   number of rule premises will be taken into account here.
   804 *}
   805 
   806 
   807 section {* Coherent Logic *}
   808 
   809 text {*
   810   \begin{matharray}{rcl}
   811     @{method_def (HOL) "coherent"} & : & @{text method} \\
   812   \end{matharray}
   813 
   814   \begin{rail}
   815     'coherent' thmrefs?
   816     ;
   817   \end{rail}
   818 
   819   The @{method (HOL) coherent} method solves problems of
   820   \emph{Coherent Logic} \cite{Bezem-Coquand:2005}, which covers
   821   applications in confluence theory, lattice theory and projective
   822   geometry.  See @{"file" "~~/src/HOL/ex/Coherent.thy"} for some
   823   examples.
   824 *}
   825 
   826 
   827 section {* Invoking automated reasoning tools -- The Sledgehammer *}
   828 
   829 text {*
   830   Isabelle/HOL includes a generic \emph{ATP manager} that allows
   831   external automated reasoning tools to crunch a pending goal.
   832   Supported provers include E\footnote{\url{http://www.eprover.org}},
   833   SPASS\footnote{\url{http://www.spass-prover.org/}}, and Vampire.
   834   There is also a wrapper to invoke provers remotely via the
   835   SystemOnTPTP\footnote{\url{http://www.cs.miami.edu/~tptp/cgi-bin/SystemOnTPTP}}
   836   web service.
   837 
   838   The problem passed to external provers consists of the goal together
   839   with a smart selection of lemmas from the current theory context.
   840   The result of a successful proof search is some source text that
   841   usually reconstructs the proof within Isabelle, without requiring
   842   external provers again.  The Metis
   843   prover\footnote{\url{http://www.gilith.com/software/metis/}} that is
   844   integrated into Isabelle/HOL is being used here.
   845 
   846   In this mode of operation, heavy means of automated reasoning are
   847   used as a strong relevance filter, while the main proof checking
   848   works via explicit inferences going through the Isabelle kernel.
   849   Moreover, rechecking Isabelle proof texts with already specified
   850   auxiliary facts is much faster than performing fully automated
   851   search over and over again.
   852 
   853   \begin{matharray}{rcl}
   854     @{command_def (HOL) "sledgehammer"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
   855     @{command_def (HOL) "print_atps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   856     @{command_def (HOL) "atp_info"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
   857     @{command_def (HOL) "atp_kill"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
   858     @{command_def (HOL) "atp_messages"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
   859     @{method_def (HOL) metis} & : & @{text method} \\
   860   \end{matharray}
   861 
   862   \begin{rail}
   863   'sledgehammer' (nameref *)
   864   ;
   865   'atp\_messages' ('(' nat ')')?
   866   ;
   867 
   868   'metis' thmrefs
   869   ;
   870   \end{rail}
   871 
   872   \begin{description}
   873 
   874   \item @{command (HOL) sledgehammer}~@{text "prover\<^sub>1 \<dots> prover\<^sub>n"}
   875   invokes the specified automated theorem provers on the first
   876   subgoal.  Provers are run in parallel, the first successful result
   877   is displayed, and the other attempts are terminated.
   878 
   879   Provers are defined in the theory context, see also @{command (HOL)
   880   print_atps}.  If no provers are given as arguments to @{command
   881   (HOL) sledgehammer}, the system refers to the default defined as
   882   ``ATP provers'' preference by the user interface.
   883 
   884   There are additional preferences for timeout (default: 60 seconds),
   885   and the maximum number of independent prover processes (default: 5);
   886   excessive provers are automatically terminated.
   887 
   888   \item @{command (HOL) print_atps} prints the list of automated
   889   theorem provers available to the @{command (HOL) sledgehammer}
   890   command.
   891 
   892   \item @{command (HOL) atp_info} prints information about presently
   893   running provers, including elapsed runtime, and the remaining time
   894   until timeout.
   895 
   896   \item @{command (HOL) atp_kill} terminates all presently running
   897   provers.
   898 
   899   \item @{command (HOL) atp_messages} displays recent messages issued
   900   by automated theorem provers.  This allows to examine results that
   901   might have got lost due to the asynchronous nature of default
   902   @{command (HOL) sledgehammer} output.  An optional message limit may
   903   be specified (default 5).
   904 
   905   \item @{method (HOL) metis}~@{text "facts"} invokes the Metis prover
   906   with the given facts.  Metis is an automated proof tool of medium
   907   strength, but is fully integrated into Isabelle/HOL, with explicit
   908   inferences going through the kernel.  Thus its results are
   909   guaranteed to be ``correct by construction''.
   910 
   911   Note that all facts used with Metis need to be specified as explicit
   912   arguments.  There are no rule declarations as for other Isabelle
   913   provers, like @{method blast} or @{method fast}.
   914 
   915   \end{description}
   916 *}
   917 
   918 
   919 section {* Unstructured case analysis and induction \label{sec:hol-induct-tac} *}
   920 
   921 text {*
   922   The following tools of Isabelle/HOL support cases analysis and
   923   induction in unstructured tactic scripts; see also
   924   \secref{sec:cases-induct} for proper Isar versions of similar ideas.
   925 
   926   \begin{matharray}{rcl}
   927     @{method_def (HOL) case_tac}@{text "\<^sup>*"} & : & @{text method} \\
   928     @{method_def (HOL) induct_tac}@{text "\<^sup>*"} & : & @{text method} \\
   929     @{method_def (HOL) ind_cases}@{text "\<^sup>*"} & : & @{text method} \\
   930     @{command_def (HOL) "inductive_cases"}@{text "\<^sup>*"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   931   \end{matharray}
   932 
   933   \begin{rail}
   934     'case\_tac' goalspec? term rule?
   935     ;
   936     'induct\_tac' goalspec? (insts * 'and') rule?
   937     ;
   938     'ind\_cases' (prop +) ('for' (name +)) ?
   939     ;
   940     'inductive\_cases' (thmdecl? (prop +) + 'and')
   941     ;
   942 
   943     rule: ('rule' ':' thmref)
   944     ;
   945   \end{rail}
   946 
   947   \begin{description}
   948 
   949   \item @{method (HOL) case_tac} and @{method (HOL) induct_tac} admit
   950   to reason about inductive types.  Rules are selected according to
   951   the declarations by the @{attribute cases} and @{attribute induct}
   952   attributes, cf.\ \secref{sec:cases-induct}.  The @{command (HOL)
   953   datatype} package already takes care of this.
   954 
   955   These unstructured tactics feature both goal addressing and dynamic
   956   instantiation.  Note that named rule cases are \emph{not} provided
   957   as would be by the proper @{method cases} and @{method induct} proof
   958   methods (see \secref{sec:cases-induct}).  Unlike the @{method
   959   induct} method, @{method induct_tac} does not handle structured rule
   960   statements, only the compact object-logic conclusion of the subgoal
   961   being addressed.
   962   
   963   \item @{method (HOL) ind_cases} and @{command (HOL)
   964   "inductive_cases"} provide an interface to the internal @{ML_text
   965   mk_cases} operation.  Rules are simplified in an unrestricted
   966   forward manner.
   967 
   968   While @{method (HOL) ind_cases} is a proof method to apply the
   969   result immediately as elimination rules, @{command (HOL)
   970   "inductive_cases"} provides case split theorems at the theory level
   971   for later use.  The @{keyword "for"} argument of the @{method (HOL)
   972   ind_cases} method allows to specify a list of variables that should
   973   be generalized before applying the resulting rule.
   974 
   975   \end{description}
   976 *}
   977 
   978 
   979 section {* Executable code *}
   980 
   981 text {*
   982   Isabelle/Pure provides two generic frameworks to support code
   983   generation from executable specifications.  Isabelle/HOL
   984   instantiates these mechanisms in a way that is amenable to end-user
   985   applications.
   986 
   987   One framework generates code from both functional and relational
   988   programs to SML.  See \cite{isabelle-HOL} for further information
   989   (this actually covers the new-style theory format as well).
   990 
   991   \begin{matharray}{rcl}
   992     @{command_def (HOL) "value"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   993     @{command_def (HOL) "code_module"} & : & @{text "theory \<rightarrow> theory"} \\
   994     @{command_def (HOL) "code_library"} & : & @{text "theory \<rightarrow> theory"} \\
   995     @{command_def (HOL) "consts_code"} & : & @{text "theory \<rightarrow> theory"} \\
   996     @{command_def (HOL) "types_code"} & : & @{text "theory \<rightarrow> theory"} \\  
   997     @{attribute_def (HOL) code} & : & @{text attribute} \\
   998   \end{matharray}
   999 
  1000   \begin{rail}
  1001   'value' term
  1002   ;
  1003 
  1004   ( 'code\_module' | 'code\_library' ) modespec ? name ? \\
  1005     ( 'file' name ) ? ( 'imports' ( name + ) ) ? \\
  1006     'contains' ( ( name '=' term ) + | term + )
  1007   ;
  1008 
  1009   modespec: '(' ( name * ) ')'
  1010   ;
  1011 
  1012   'consts\_code' (codespec +)
  1013   ;
  1014 
  1015   codespec: const template attachment ?
  1016   ;
  1017 
  1018   'types\_code' (tycodespec +)
  1019   ;
  1020 
  1021   tycodespec: name template attachment ?
  1022   ;
  1023 
  1024   const: term
  1025   ;
  1026 
  1027   template: '(' string ')'
  1028   ;
  1029 
  1030   attachment: 'attach' modespec ? verblbrace text verbrbrace
  1031   ;
  1032 
  1033   'code' (name)?
  1034   ;
  1035   \end{rail}
  1036 
  1037   \begin{description}
  1038 
  1039   \item @{command (HOL) "value"}~@{text t} evaluates and prints a term
  1040   using the code generator.
  1041 
  1042   \end{description}
  1043 
  1044   \medskip The other framework generates code from functional programs
  1045   (including overloading using type classes) to SML \cite{SML}, OCaml
  1046   \cite{OCaml} and Haskell \cite{haskell-revised-report}.
  1047   Conceptually, code generation is split up in three steps:
  1048   \emph{selection} of code theorems, \emph{translation} into an
  1049   abstract executable view and \emph{serialization} to a specific
  1050   \emph{target language}.  See \cite{isabelle-codegen} for an
  1051   introduction on how to use it.
  1052 
  1053   \begin{matharray}{rcl}
  1054     @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1055     @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1056     @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1057     @{command_def (HOL) "code_datatype"} & : & @{text "theory \<rightarrow> theory"} \\
  1058     @{command_def (HOL) "code_const"} & : & @{text "theory \<rightarrow> theory"} \\
  1059     @{command_def (HOL) "code_type"} & : & @{text "theory \<rightarrow> theory"} \\
  1060     @{command_def (HOL) "code_class"} & : & @{text "theory \<rightarrow> theory"} \\
  1061     @{command_def (HOL) "code_instance"} & : & @{text "theory \<rightarrow> theory"} \\
  1062     @{command_def (HOL) "code_monad"} & : & @{text "theory \<rightarrow> theory"} \\
  1063     @{command_def (HOL) "code_reserved"} & : & @{text "theory \<rightarrow> theory"} \\
  1064     @{command_def (HOL) "code_include"} & : & @{text "theory \<rightarrow> theory"} \\
  1065     @{command_def (HOL) "code_modulename"} & : & @{text "theory \<rightarrow> theory"} \\
  1066     @{command_def (HOL) "code_abort"} & : & @{text "theory \<rightarrow> theory"} \\
  1067     @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1068     @{attribute_def (HOL) code} & : & @{text attribute} \\
  1069   \end{matharray}
  1070 
  1071   \begin{rail}
  1072     'export\_code' ( constexpr + ) ? \\
  1073       ( ( 'in' target ( 'module\_name' string ) ? \\
  1074         ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
  1075     ;
  1076 
  1077     'code\_thms' ( constexpr + ) ?
  1078     ;
  1079 
  1080     'code\_deps' ( constexpr + ) ?
  1081     ;
  1082 
  1083     const: term
  1084     ;
  1085 
  1086     constexpr: ( const | 'name.*' | '*' )
  1087     ;
  1088 
  1089     typeconstructor: nameref
  1090     ;
  1091 
  1092     class: nameref
  1093     ;
  1094 
  1095     target: 'OCaml' | 'SML' | 'Haskell'
  1096     ;
  1097 
  1098     'code\_datatype' const +
  1099     ;
  1100 
  1101     'code\_const' (const + 'and') \\
  1102       ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
  1103     ;
  1104 
  1105     'code\_type' (typeconstructor + 'and') \\
  1106       ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
  1107     ;
  1108 
  1109     'code\_class' (class + 'and') \\
  1110       ( ( '(' target \\ ( string ? + 'and' ) ')' ) + )
  1111     ;
  1112 
  1113     'code\_instance' (( typeconstructor '::' class ) + 'and') \\
  1114       ( ( '(' target ( '-' ? + 'and' ) ')' ) + )
  1115     ;
  1116 
  1117     'code\_monad' const const target
  1118     ;
  1119 
  1120     'code\_reserved' target ( string + )
  1121     ;
  1122 
  1123     'code\_include' target ( string ( string | '-') )
  1124     ;
  1125 
  1126     'code\_modulename' target ( ( string string ) + )
  1127     ;
  1128 
  1129     'code\_abort' ( const + )
  1130     ;
  1131 
  1132     syntax: string | ( 'infix' | 'infixl' | 'infixr' ) nat string
  1133     ;
  1134 
  1135     'code' ( 'inline' ) ? ( 'del' ) ?
  1136     ;
  1137   \end{rail}
  1138 
  1139   \begin{description}
  1140 
  1141   \item @{command (HOL) "export_code"} is the canonical interface for
  1142   generating and serializing code: for a given list of constants, code
  1143   is generated for the specified target languages.  Abstract code is
  1144   cached incrementally.  If no constant is given, the currently cached
  1145   code is serialized.  If no serialization instruction is given, only
  1146   abstract code is cached.
  1147 
  1148   Constants may be specified by giving them literally, referring to
  1149   all executable contants within a certain theory by giving @{text
  1150   "name.*"}, or referring to \emph{all} executable constants currently
  1151   available by giving @{text "*"}.
  1152 
  1153   By default, for each involved theory one corresponding name space
  1154   module is generated.  Alternativly, a module name may be specified
  1155   after the @{keyword "module_name"} keyword; then \emph{all} code is
  1156   placed in this module.
  1157 
  1158   For \emph{SML} and \emph{OCaml}, the file specification refers to a
  1159   single file; for \emph{Haskell}, it refers to a whole directory,
  1160   where code is generated in multiple files reflecting the module
  1161   hierarchy.  The file specification ``@{text "-"}'' denotes standard
  1162   output.  For \emph{SML}, omitting the file specification compiles
  1163   code internally in the context of the current ML session.
  1164 
  1165   Serializers take an optional list of arguments in parentheses.  For
  1166   \emph{Haskell} a module name prefix may be given using the ``@{text
  1167   "root:"}'' argument; ``@{text string_classes}'' adds a ``@{verbatim
  1168   "deriving (Read, Show)"}'' clause to each appropriate datatype
  1169   declaration.
  1170 
  1171   \item @{command (HOL) "code_thms"} prints a list of theorems
  1172   representing the corresponding program containing all given
  1173   constants; if no constants are given, the currently cached code
  1174   theorems are printed.
  1175 
  1176   \item @{command (HOL) "code_deps"} visualizes dependencies of
  1177   theorems representing the corresponding program containing all given
  1178   constants; if no constants are given, the currently cached code
  1179   theorems are visualized.
  1180 
  1181   \item @{command (HOL) "code_datatype"} specifies a constructor set
  1182   for a logical type.
  1183 
  1184   \item @{command (HOL) "code_const"} associates a list of constants
  1185   with target-specific serializations; omitting a serialization
  1186   deletes an existing serialization.
  1187 
  1188   \item @{command (HOL) "code_type"} associates a list of type
  1189   constructors with target-specific serializations; omitting a
  1190   serialization deletes an existing serialization.
  1191 
  1192   \item @{command (HOL) "code_class"} associates a list of classes
  1193   with target-specific class names; omitting a serialization deletes
  1194   an existing serialization.  This applies only to \emph{Haskell}.
  1195 
  1196   \item @{command (HOL) "code_instance"} declares a list of type
  1197   constructor / class instance relations as ``already present'' for a
  1198   given target.  Omitting a ``@{text "-"}'' deletes an existing
  1199   ``already present'' declaration.  This applies only to
  1200   \emph{Haskell}.
  1201 
  1202   \item @{command (HOL) "code_monad"} provides an auxiliary mechanism
  1203   to generate monadic code for Haskell.
  1204 
  1205   \item @{command (HOL) "code_reserved"} declares a list of names as
  1206   reserved for a given target, preventing it to be shadowed by any
  1207   generated code.
  1208 
  1209   \item @{command (HOL) "code_include"} adds arbitrary named content
  1210   (``include'') to generated code.  A ``@{text "-"}'' as last argument
  1211   will remove an already added ``include''.
  1212 
  1213   \item @{command (HOL) "code_modulename"} declares aliasings from one
  1214   module name onto another.
  1215 
  1216   \item @{command (HOL) "code_abort"} declares constants which are not
  1217   required to have a definition by means of code equations; if
  1218   needed these are implemented by program abort instead.
  1219 
  1220   \item @{attribute (HOL) code} explicitly selects (or with option
  1221   ``@{text "del"}'' deselects) a code equation for code
  1222   generation.  Usually packages introducing code equations provide
  1223   a reasonable default setup for selection.
  1224 
  1225   \item @{attribute (HOL) code}~@{text inline} declares (or with
  1226   option ``@{text "del"}'' removes) inlining theorems which are
  1227   applied as rewrite rules to any code equation during
  1228   preprocessing.
  1229 
  1230   \item @{command (HOL) "print_codesetup"} gives an overview on
  1231   selected code equations, code generator datatypes and
  1232   preprocessor setup.
  1233 
  1234   \end{description}
  1235 *}
  1236 
  1237 
  1238 section {* Definition by specification \label{sec:hol-specification} *}
  1239 
  1240 text {*
  1241   \begin{matharray}{rcl}
  1242     @{command_def (HOL) "specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
  1243     @{command_def (HOL) "ax_specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
  1244   \end{matharray}
  1245 
  1246   \begin{rail}
  1247   ('specification' | 'ax\_specification') '(' (decl +) ')' \\ (thmdecl? prop +)
  1248   ;
  1249   decl: ((name ':')? term '(' 'overloaded' ')'?)
  1250   \end{rail}
  1251 
  1252   \begin{description}
  1253 
  1254   \item @{command (HOL) "specification"}~@{text "decls \<phi>"} sets up a
  1255   goal stating the existence of terms with the properties specified to
  1256   hold for the constants given in @{text decls}.  After finishing the
  1257   proof, the theory will be augmented with definitions for the given
  1258   constants, as well as with theorems stating the properties for these
  1259   constants.
  1260 
  1261   \item @{command (HOL) "ax_specification"}~@{text "decls \<phi>"} sets up
  1262   a goal stating the existence of terms with the properties specified
  1263   to hold for the constants given in @{text decls}.  After finishing
  1264   the proof, the theory will be augmented with axioms expressing the
  1265   properties given in the first place.
  1266 
  1267   \item @{text decl} declares a constant to be defined by the
  1268   specification given.  The definition for the constant @{text c} is
  1269   bound to the name @{text c_def} unless a theorem name is given in
  1270   the declaration.  Overloaded constants should be declared as such.
  1271 
  1272   \end{description}
  1273 
  1274   Whether to use @{command (HOL) "specification"} or @{command (HOL)
  1275   "ax_specification"} is to some extent a matter of style.  @{command
  1276   (HOL) "specification"} introduces no new axioms, and so by
  1277   construction cannot introduce inconsistencies, whereas @{command
  1278   (HOL) "ax_specification"} does introduce axioms, but only after the
  1279   user has explicitly proven it to be safe.  A practical issue must be
  1280   considered, though: After introducing two constants with the same
  1281   properties using @{command (HOL) "specification"}, one can prove
  1282   that the two constants are, in fact, equal.  If this might be a
  1283   problem, one should use @{command (HOL) "ax_specification"}.
  1284 *}
  1285 
  1286 end