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