doc-src/IsarRef/Thy/HOL_Specific.thy
author blanchet
Fri, 27 May 2011 10:30:08 +0200
changeset 43860 619f16bf2150
parent 43857 42330f25142c
child 43881 665623e695ea
permissions -rw-r--r--
removed FIXME: "try_methods" doesn't take a clasimpmod, because it needs to generate Isar proof text in case of success and hence needs total control over its arguments
     1 theory HOL_Specific
     2 imports Base Main
     3 begin
     4 
     5 chapter {* Isabelle/HOL \label{ch:hol} *}
     6 
     7 section {* Typedef axiomatization \label{sec:hol-typedef} *}
     8 
     9 text {*
    10   \begin{matharray}{rcl}
    11     @{command_def (HOL) "typedef"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
    12   \end{matharray}
    13 
    14   @{rail "
    15     @@{command (HOL) typedef} altname? abstype '=' repset
    16     ;
    17 
    18     altname: '(' (@{syntax name} | @'open' | @'open' @{syntax name}) ')'
    19     ;
    20     abstype: @{syntax typespec_sorts} @{syntax mixfix}?
    21     ;
    22     repset: @{syntax term} (@'morphisms' @{syntax name} @{syntax name})?
    23   "}
    24 
    25   \begin{description}
    26 
    27   \item @{command (HOL) "typedef"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = A"}
    28   axiomatizes a Gordon/HOL-style type definition in the background
    29   theory of the current context, depending on a non-emptiness result
    30   of the set @{text A} (which needs to be proven interactively).
    31 
    32   The raw type may not depend on parameters or assumptions of the
    33   context --- this is logically impossible in Isabelle/HOL --- but the
    34   non-emptiness property can be local, potentially resulting in
    35   multiple interpretations in target contexts.  Thus the established
    36   bijection between the representing set @{text A} and the new type
    37   @{text t} may semantically depend on local assumptions.
    38 
    39   By default, @{command (HOL) "typedef"} defines both a type @{text t}
    40   and a set (term constant) of the same name, unless an alternative
    41   base name is given in parentheses, or the ``@{text "(open)"}''
    42   declaration is used to suppress a separate constant definition
    43   altogether.  The injection from type to set is called @{text Rep_t},
    44   its inverse @{text Abs_t} --- this may be changed via an explicit
    45   @{keyword (HOL) "morphisms"} declaration.
    46 
    47   Theorems @{text Rep_t}, @{text Rep_t_inverse}, and @{text
    48   Abs_t_inverse} provide the most basic characterization as a
    49   corresponding injection/surjection pair (in both directions).  Rules
    50   @{text Rep_t_inject} and @{text Abs_t_inject} provide a slightly
    51   more convenient view on the injectivity part, suitable for automated
    52   proof tools (e.g.\ in @{attribute simp} or @{attribute iff}
    53   declarations).  Rules @{text Rep_t_cases}/@{text Rep_t_induct}, and
    54   @{text Abs_t_cases}/@{text Abs_t_induct} provide alternative views
    55   on surjectivity; these are already declared as set or type rules for
    56   the generic @{method cases} and @{method induct} methods.
    57 
    58   An alternative name for the set definition (and other derived
    59   entities) may be specified in parentheses; the default is to use
    60   @{text t} as indicated before.
    61 
    62   \end{description}
    63 *}
    64 
    65 
    66 section {* Adhoc tuples *}
    67 
    68 text {*
    69   \begin{matharray}{rcl}
    70     @{attribute_def (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\
    71   \end{matharray}
    72 
    73   @{rail "
    74     @@{attribute (HOL) split_format} ('(' 'complete' ')')?
    75   "}
    76 
    77   \begin{description}
    78 
    79   \item @{attribute (HOL) split_format}\ @{text "(complete)"} causes
    80   arguments in function applications to be represented canonically
    81   according to their tuple type structure.
    82 
    83   Note that this operation tends to invent funny names for new local
    84   parameters introduced.
    85 
    86   \end{description}
    87 *}
    88 
    89 
    90 section {* Records \label{sec:hol-record} *}
    91 
    92 text {*
    93   In principle, records merely generalize the concept of tuples, where
    94   components may be addressed by labels instead of just position.  The
    95   logical infrastructure of records in Isabelle/HOL is slightly more
    96   advanced, though, supporting truly extensible record schemes.  This
    97   admits operations that are polymorphic with respect to record
    98   extension, yielding ``object-oriented'' effects like (single)
    99   inheritance.  See also \cite{NaraschewskiW-TPHOLs98} for more
   100   details on object-oriented verification and record subtyping in HOL.
   101 *}
   102 
   103 
   104 subsection {* Basic concepts *}
   105 
   106 text {*
   107   Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records
   108   at the level of terms and types.  The notation is as follows:
   109 
   110   \begin{center}
   111   \begin{tabular}{l|l|l}
   112     & record terms & record types \\ \hline
   113     fixed & @{text "\<lparr>x = a, y = b\<rparr>"} & @{text "\<lparr>x :: A, y :: B\<rparr>"} \\
   114     schematic & @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} &
   115       @{text "\<lparr>x :: A, y :: B, \<dots> :: M\<rparr>"} \\
   116   \end{tabular}
   117   \end{center}
   118 
   119   \noindent The ASCII representation of @{text "\<lparr>x = a\<rparr>"} is @{text
   120   "(| x = a |)"}.
   121 
   122   A fixed record @{text "\<lparr>x = a, y = b\<rparr>"} has field @{text x} of value
   123   @{text a} and field @{text y} of value @{text b}.  The corresponding
   124   type is @{text "\<lparr>x :: A, y :: B\<rparr>"}, assuming that @{text "a :: A"}
   125   and @{text "b :: B"}.
   126 
   127   A record scheme like @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} contains fields
   128   @{text x} and @{text y} as before, but also possibly further fields
   129   as indicated by the ``@{text "\<dots>"}'' notation (which is actually part
   130   of the syntax).  The improper field ``@{text "\<dots>"}'' of a record
   131   scheme is called the \emph{more part}.  Logically it is just a free
   132   variable, which is occasionally referred to as ``row variable'' in
   133   the literature.  The more part of a record scheme may be
   134   instantiated by zero or more further components.  For example, the
   135   previous scheme may get instantiated to @{text "\<lparr>x = a, y = b, z =
   136   c, \<dots> = m'\<rparr>"}, where @{text m'} refers to a different more part.
   137   Fixed records are special instances of record schemes, where
   138   ``@{text "\<dots>"}'' is properly terminated by the @{text "() :: unit"}
   139   element.  In fact, @{text "\<lparr>x = a, y = b\<rparr>"} is just an abbreviation
   140   for @{text "\<lparr>x = a, y = b, \<dots> = ()\<rparr>"}.
   141 
   142   \medskip Two key observations make extensible records in a simply
   143   typed language like HOL work out:
   144 
   145   \begin{enumerate}
   146 
   147   \item the more part is internalized, as a free term or type
   148   variable,
   149 
   150   \item field names are externalized, they cannot be accessed within
   151   the logic as first-class values.
   152 
   153   \end{enumerate}
   154 
   155   \medskip In Isabelle/HOL record types have to be defined explicitly,
   156   fixing their field names and types, and their (optional) parent
   157   record.  Afterwards, records may be formed using above syntax, while
   158   obeying the canonical order of fields as given by their declaration.
   159   The record package provides several standard operations like
   160   selectors and updates.  The common setup for various generic proof
   161   tools enable succinct reasoning patterns.  See also the Isabelle/HOL
   162   tutorial \cite{isabelle-hol-book} for further instructions on using
   163   records in practice.
   164 *}
   165 
   166 
   167 subsection {* Record specifications *}
   168 
   169 text {*
   170   \begin{matharray}{rcl}
   171     @{command_def (HOL) "record"} & : & @{text "theory \<rightarrow> theory"} \\
   172   \end{matharray}
   173 
   174   @{rail "
   175     @@{command (HOL) record} @{syntax typespec_sorts} '=' \\
   176       (@{syntax type} '+')? (@{syntax constdecl} +)
   177   "}
   178 
   179   \begin{description}
   180 
   181   \item @{command (HOL) "record"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t = \<tau> + c\<^sub>1 :: \<sigma>\<^sub>1
   182   \<dots> c\<^sub>n :: \<sigma>\<^sub>n"} defines extensible record type @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"},
   183   derived from the optional parent record @{text "\<tau>"} by adding new
   184   field components @{text "c\<^sub>i :: \<sigma>\<^sub>i"} etc.
   185 
   186   The type variables of @{text "\<tau>"} and @{text "\<sigma>\<^sub>i"} need to be
   187   covered by the (distinct) parameters @{text "\<alpha>\<^sub>1, \<dots>,
   188   \<alpha>\<^sub>m"}.  Type constructor @{text t} has to be new, while @{text
   189   \<tau>} needs to specify an instance of an existing record type.  At
   190   least one new field @{text "c\<^sub>i"} has to be specified.
   191   Basically, field names need to belong to a unique record.  This is
   192   not a real restriction in practice, since fields are qualified by
   193   the record name internally.
   194 
   195   The parent record specification @{text \<tau>} is optional; if omitted
   196   @{text t} becomes a root record.  The hierarchy of all records
   197   declared within a theory context forms a forest structure, i.e.\ a
   198   set of trees starting with a root record each.  There is no way to
   199   merge multiple parent records!
   200 
   201   For convenience, @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} is made a
   202   type abbreviation for the fixed record type @{text "\<lparr>c\<^sub>1 ::
   203   \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n\<rparr>"}, likewise is @{text
   204   "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m, \<zeta>) t_scheme"} made an abbreviation for
   205   @{text "\<lparr>c\<^sub>1 :: \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n, \<dots> ::
   206   \<zeta>\<rparr>"}.
   207 
   208   \end{description}
   209 *}
   210 
   211 
   212 subsection {* Record operations *}
   213 
   214 text {*
   215   Any record definition of the form presented above produces certain
   216   standard operations.  Selectors and updates are provided for any
   217   field, including the improper one ``@{text more}''.  There are also
   218   cumulative record constructor functions.  To simplify the
   219   presentation below, we assume for now that @{text "(\<alpha>\<^sub>1, \<dots>,
   220   \<alpha>\<^sub>m) t"} is a root record with fields @{text "c\<^sub>1 ::
   221   \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n"}.
   222 
   223   \medskip \textbf{Selectors} and \textbf{updates} are available for
   224   any field (including ``@{text more}''):
   225 
   226   \begin{matharray}{lll}
   227     @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
   228     @{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>"} \\
   229   \end{matharray}
   230 
   231   There is special syntax for application of updates: @{text "r\<lparr>x :=
   232   a\<rparr>"} abbreviates term @{text "x_update a r"}.  Further notation for
   233   repeated updates is also available: @{text "r\<lparr>x := a\<rparr>\<lparr>y := b\<rparr>\<lparr>z :=
   234   c\<rparr>"} may be written @{text "r\<lparr>x := a, y := b, z := c\<rparr>"}.  Note that
   235   because of postfix notation the order of fields shown here is
   236   reverse than in the actual term.  Since repeated updates are just
   237   function applications, fields may be freely permuted in @{text "\<lparr>x
   238   := a, y := b, z := c\<rparr>"}, as far as logical equality is concerned.
   239   Thus commutativity of independent updates can be proven within the
   240   logic for any two fields, but not as a general theorem.
   241 
   242   \medskip The \textbf{make} operation provides a cumulative record
   243   constructor function:
   244 
   245   \begin{matharray}{lll}
   246     @{text "t.make"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   247   \end{matharray}
   248 
   249   \medskip We now reconsider the case of non-root records, which are
   250   derived of some parent.  In general, the latter may depend on
   251   another parent as well, resulting in a list of \emph{ancestor
   252   records}.  Appending the lists of fields of all ancestors results in
   253   a certain field prefix.  The record package automatically takes care
   254   of this by lifting operations over this context of ancestor fields.
   255   Assuming that @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} has ancestor
   256   fields @{text "b\<^sub>1 :: \<rho>\<^sub>1, \<dots>, b\<^sub>k :: \<rho>\<^sub>k"},
   257   the above record operations will get the following types:
   258 
   259   \medskip
   260   \begin{tabular}{lll}
   261     @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
   262     @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow>
   263       \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow>
   264       \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
   265     @{text "t.make"} & @{text "::"} & @{text "\<rho>\<^sub>1 \<Rightarrow> \<dots> \<rho>\<^sub>k \<Rightarrow> \<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow>
   266       \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   267   \end{tabular}
   268   \medskip
   269 
   270   \noindent Some further operations address the extension aspect of a
   271   derived record scheme specifically: @{text "t.fields"} produces a
   272   record fragment consisting of exactly the new fields introduced here
   273   (the result may serve as a more part elsewhere); @{text "t.extend"}
   274   takes a fixed record and adds a given more part; @{text
   275   "t.truncate"} restricts a record scheme to a fixed record.
   276 
   277   \medskip
   278   \begin{tabular}{lll}
   279     @{text "t.fields"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   280     @{text "t.extend"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr> \<Rightarrow>
   281       \<zeta> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
   282     @{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>"} \\
   283   \end{tabular}
   284   \medskip
   285 
   286   \noindent Note that @{text "t.make"} and @{text "t.fields"} coincide
   287   for root records.
   288 *}
   289 
   290 
   291 subsection {* Derived rules and proof tools *}
   292 
   293 text {*
   294   The record package proves several results internally, declaring
   295   these facts to appropriate proof tools.  This enables users to
   296   reason about record structures quite conveniently.  Assume that
   297   @{text t} is a record type as specified above.
   298 
   299   \begin{enumerate}
   300 
   301   \item Standard conversions for selectors or updates applied to
   302   record constructor terms are made part of the default Simplifier
   303   context; thus proofs by reduction of basic operations merely require
   304   the @{method simp} method without further arguments.  These rules
   305   are available as @{text "t.simps"}, too.
   306 
   307   \item Selectors applied to updated records are automatically reduced
   308   by an internal simplification procedure, which is also part of the
   309   standard Simplifier setup.
   310 
   311   \item Inject equations of a form analogous to @{prop "(x, y) = (x',
   312   y') \<equiv> x = x' \<and> y = y'"} are declared to the Simplifier and Classical
   313   Reasoner as @{attribute iff} rules.  These rules are available as
   314   @{text "t.iffs"}.
   315 
   316   \item The introduction rule for record equality analogous to @{text
   317   "x r = x r' \<Longrightarrow> y r = y r' \<dots> \<Longrightarrow> r = r'"} is declared to the Simplifier,
   318   and as the basic rule context as ``@{attribute intro}@{text "?"}''.
   319   The rule is called @{text "t.equality"}.
   320 
   321   \item Representations of arbitrary record expressions as canonical
   322   constructor terms are provided both in @{method cases} and @{method
   323   induct} format (cf.\ the generic proof methods of the same name,
   324   \secref{sec:cases-induct}).  Several variations are available, for
   325   fixed records, record schemes, more parts etc.
   326 
   327   The generic proof methods are sufficiently smart to pick the most
   328   sensible rule according to the type of the indicated record
   329   expression: users just need to apply something like ``@{text "(cases
   330   r)"}'' to a certain proof problem.
   331 
   332   \item The derived record operations @{text "t.make"}, @{text
   333   "t.fields"}, @{text "t.extend"}, @{text "t.truncate"} are \emph{not}
   334   treated automatically, but usually need to be expanded by hand,
   335   using the collective fact @{text "t.defs"}.
   336 
   337   \end{enumerate}
   338 *}
   339 
   340 
   341 section {* Datatypes \label{sec:hol-datatype} *}
   342 
   343 text {*
   344   \begin{matharray}{rcl}
   345     @{command_def (HOL) "datatype"} & : & @{text "theory \<rightarrow> theory"} \\
   346     @{command_def (HOL) "rep_datatype"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   347   \end{matharray}
   348 
   349   @{rail "
   350     @@{command (HOL) datatype} (spec + @'and')
   351     ;
   352     @@{command (HOL) rep_datatype} ('(' (@{syntax name} +) ')')? (@{syntax term} +)
   353     ;
   354 
   355     spec: @{syntax parname}? @{syntax typespec} @{syntax mixfix}? '=' (cons + '|')
   356     ;
   357     cons: @{syntax name} (@{syntax type} * ) @{syntax mixfix}?
   358   "}
   359 
   360   \begin{description}
   361 
   362   \item @{command (HOL) "datatype"} defines inductive datatypes in
   363   HOL.
   364 
   365   \item @{command (HOL) "rep_datatype"} represents existing types as
   366   inductive ones, generating the standard infrastructure of derived
   367   concepts (primitive recursion etc.).
   368 
   369   \end{description}
   370 
   371   The induction and exhaustion theorems generated provide case names
   372   according to the constructors involved, while parameters are named
   373   after the types (see also \secref{sec:cases-induct}).
   374 
   375   See \cite{isabelle-HOL} for more details on datatypes, but beware of
   376   the old-style theory syntax being used there!  Apart from proper
   377   proof methods for case-analysis and induction, there are also
   378   emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL)
   379   induct_tac} available, see \secref{sec:hol-induct-tac}; these admit
   380   to refer directly to the internal structure of subgoals (including
   381   internally bound parameters).
   382 *}
   383 
   384 
   385 section {* Functorial structure of types *}
   386 
   387 text {*
   388   \begin{matharray}{rcl}
   389     @{command_def (HOL) "enriched_type"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
   390   \end{matharray}
   391 
   392   @{rail "
   393     @@{command (HOL) enriched_type} (@{syntax name} ':')? @{syntax term}
   394     ;
   395   "}
   396 
   397   \begin{description}
   398 
   399   \item @{command (HOL) "enriched_type"}~@{text "prefix: m"} allows to
   400   prove and register properties about the functorial structure of type
   401   constructors.  These properties then can be used by other packages
   402   to deal with those type constructors in certain type constructions.
   403   Characteristic theorems are noted in the current local theory.  By
   404   default, they are prefixed with the base name of the type
   405   constructor, an explicit prefix can be given alternatively.
   406 
   407   The given term @{text "m"} is considered as \emph{mapper} for the
   408   corresponding type constructor and must conform to the following
   409   type pattern:
   410 
   411   \begin{matharray}{lll}
   412     @{text "m"} & @{text "::"} &
   413       @{text "\<sigma>\<^isub>1 \<Rightarrow> \<dots> \<sigma>\<^isub>k \<Rightarrow> (\<^vec>\<alpha>\<^isub>n) t \<Rightarrow> (\<^vec>\<beta>\<^isub>n) t"} \\
   414   \end{matharray}
   415 
   416   \noindent where @{text t} is the type constructor, @{text
   417   "\<^vec>\<alpha>\<^isub>n"} and @{text "\<^vec>\<beta>\<^isub>n"} are distinct
   418   type variables free in the local theory and @{text "\<sigma>\<^isub>1"},
   419   \ldots, @{text "\<sigma>\<^isub>k"} is a subsequence of @{text "\<alpha>\<^isub>1 \<Rightarrow>
   420   \<beta>\<^isub>1"}, @{text "\<beta>\<^isub>1 \<Rightarrow> \<alpha>\<^isub>1"}, \ldots,
   421   @{text "\<alpha>\<^isub>n \<Rightarrow> \<beta>\<^isub>n"}, @{text "\<beta>\<^isub>n \<Rightarrow>
   422   \<alpha>\<^isub>n"}.
   423 
   424   \end{description}
   425 *}
   426 
   427 
   428 section {* Recursive functions \label{sec:recursion} *}
   429 
   430 text {*
   431   \begin{matharray}{rcl}
   432     @{command_def (HOL) "primrec"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   433     @{command_def (HOL) "fun"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   434     @{command_def (HOL) "function"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
   435     @{command_def (HOL) "termination"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
   436   \end{matharray}
   437 
   438   @{rail "
   439     @@{command (HOL) primrec} @{syntax target}? @{syntax \"fixes\"} @'where' equations
   440     ;
   441     (@@{command (HOL) fun} | @@{command (HOL) function}) @{syntax target}? functionopts?
   442       @{syntax \"fixes\"} \\ @'where' equations
   443     ;
   444 
   445     equations: (@{syntax thmdecl}? @{syntax prop} + '|')
   446     ;
   447     functionopts: '(' (('sequential' | 'domintros') + ',') ')'
   448     ;
   449     @@{command (HOL) termination} @{syntax term}?
   450   "}
   451 
   452   \begin{description}
   453 
   454   \item @{command (HOL) "primrec"} defines primitive recursive
   455   functions over datatypes, see also \cite{isabelle-HOL}.
   456 
   457   \item @{command (HOL) "function"} defines functions by general
   458   wellfounded recursion. A detailed description with examples can be
   459   found in \cite{isabelle-function}. The function is specified by a
   460   set of (possibly conditional) recursive equations with arbitrary
   461   pattern matching. The command generates proof obligations for the
   462   completeness and the compatibility of patterns.
   463 
   464   The defined function is considered partial, and the resulting
   465   simplification rules (named @{text "f.psimps"}) and induction rule
   466   (named @{text "f.pinduct"}) are guarded by a generated domain
   467   predicate @{text "f_dom"}. The @{command (HOL) "termination"}
   468   command can then be used to establish that the function is total.
   469 
   470   \item @{command (HOL) "fun"} is a shorthand notation for ``@{command
   471   (HOL) "function"}~@{text "(sequential)"}, followed by automated
   472   proof attempts regarding pattern matching and termination.  See
   473   \cite{isabelle-function} for further details.
   474 
   475   \item @{command (HOL) "termination"}~@{text f} commences a
   476   termination proof for the previously defined function @{text f}.  If
   477   this is omitted, the command refers to the most recent function
   478   definition.  After the proof is closed, the recursive equations and
   479   the induction principle is established.
   480 
   481   \end{description}
   482 
   483   Recursive definitions introduced by the @{command (HOL) "function"}
   484   command accommodate
   485   reasoning by induction (cf.\ \secref{sec:cases-induct}): rule @{text
   486   "c.induct"} (where @{text c} is the name of the function definition)
   487   refers to a specific induction rule, with parameters named according
   488   to the user-specified equations. Cases are numbered (starting from 1).
   489 
   490   For @{command (HOL) "primrec"}, the induction principle coincides
   491   with structural recursion on the datatype the recursion is carried
   492   out.
   493 
   494   The equations provided by these packages may be referred later as
   495   theorem list @{text "f.simps"}, where @{text f} is the (collective)
   496   name of the functions defined.  Individual equations may be named
   497   explicitly as well.
   498 
   499   The @{command (HOL) "function"} command accepts the following
   500   options.
   501 
   502   \begin{description}
   503 
   504   \item @{text sequential} enables a preprocessor which disambiguates
   505   overlapping patterns by making them mutually disjoint.  Earlier
   506   equations take precedence over later ones.  This allows to give the
   507   specification in a format very similar to functional programming.
   508   Note that the resulting simplification and induction rules
   509   correspond to the transformed specification, not the one given
   510   originally. This usually means that each equation given by the user
   511   may result in several theorems.  Also note that this automatic
   512   transformation only works for ML-style datatype patterns.
   513 
   514   \item @{text domintros} enables the automated generation of
   515   introduction rules for the domain predicate. While mostly not
   516   needed, they can be helpful in some proofs about partial functions.
   517 
   518   \end{description}
   519 *}
   520 
   521 
   522 subsection {* Proof methods related to recursive definitions *}
   523 
   524 text {*
   525   \begin{matharray}{rcl}
   526     @{method_def (HOL) pat_completeness} & : & @{text method} \\
   527     @{method_def (HOL) relation} & : & @{text method} \\
   528     @{method_def (HOL) lexicographic_order} & : & @{text method} \\
   529     @{method_def (HOL) size_change} & : & @{text method} \\
   530   \end{matharray}
   531 
   532   @{rail "
   533     @@{method (HOL) relation} @{syntax term}
   534     ;
   535     @@{method (HOL) lexicographic_order} (@{syntax clasimpmod} * )
   536     ;
   537     @@{method (HOL) size_change} ( orders (@{syntax clasimpmod} * ) )
   538     ;
   539     orders: ( 'max' | 'min' | 'ms' ) *
   540   "}
   541 
   542   \begin{description}
   543 
   544   \item @{method (HOL) pat_completeness} is a specialized method to
   545   solve goals regarding the completeness of pattern matching, as
   546   required by the @{command (HOL) "function"} package (cf.\
   547   \cite{isabelle-function}).
   548 
   549   \item @{method (HOL) relation}~@{text R} introduces a termination
   550   proof using the relation @{text R}.  The resulting proof state will
   551   contain goals expressing that @{text R} is wellfounded, and that the
   552   arguments of recursive calls decrease with respect to @{text R}.
   553   Usually, this method is used as the initial proof step of manual
   554   termination proofs.
   555 
   556   \item @{method (HOL) "lexicographic_order"} attempts a fully
   557   automated termination proof by searching for a lexicographic
   558   combination of size measures on the arguments of the function. The
   559   method accepts the same arguments as the @{method auto} method,
   560   which it uses internally to prove local descents.  The same context
   561   modifiers as for @{method auto} are accepted, see
   562   \secref{sec:clasimp}.
   563 
   564   In case of failure, extensive information is printed, which can help
   565   to analyse the situation (cf.\ \cite{isabelle-function}).
   566 
   567   \item @{method (HOL) "size_change"} also works on termination goals,
   568   using a variation of the size-change principle, together with a
   569   graph decomposition technique (see \cite{krauss_phd} for details).
   570   Three kinds of orders are used internally: @{text max}, @{text min},
   571   and @{text ms} (multiset), which is only available when the theory
   572   @{text Multiset} is loaded. When no order kinds are given, they are
   573   tried in order. The search for a termination proof uses SAT solving
   574   internally.
   575 
   576  For local descent proofs, the same context modifiers as for @{method
   577   auto} are accepted, see \secref{sec:clasimp}.
   578 
   579   \end{description}
   580 *}
   581 
   582 subsection {* Functions with explicit partiality *}
   583 
   584 text {*
   585   \begin{matharray}{rcl}
   586     @{command_def (HOL) "partial_function"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   587     @{attribute_def (HOL) "partial_function_mono"} & : & @{text attribute} \\
   588   \end{matharray}
   589 
   590   @{rail "
   591     @@{command (HOL) partial_function} @{syntax target}?
   592       '(' @{syntax nameref} ')' @{syntax \"fixes\"} \\
   593       @'where' @{syntax thmdecl}? @{syntax prop}
   594   "}
   595 
   596   \begin{description}
   597 
   598   \item @{command (HOL) "partial_function"}~@{text "(mode)"} defines
   599   recursive functions based on fixpoints in complete partial
   600   orders. No termination proof is required from the user or
   601   constructed internally. Instead, the possibility of non-termination
   602   is modelled explicitly in the result type, which contains an
   603   explicit bottom element.
   604 
   605   Pattern matching and mutual recursion are currently not supported.
   606   Thus, the specification consists of a single function described by a
   607   single recursive equation.
   608 
   609   There are no fixed syntactic restrictions on the body of the
   610   function, but the induced functional must be provably monotonic
   611   wrt.\ the underlying order.  The monotonicitity proof is performed
   612   internally, and the definition is rejected when it fails. The proof
   613   can be influenced by declaring hints using the
   614   @{attribute (HOL) partial_function_mono} attribute.
   615 
   616   The mandatory @{text mode} argument specifies the mode of operation
   617   of the command, which directly corresponds to a complete partial
   618   order on the result type. By default, the following modes are
   619   defined:
   620 
   621   \begin{description}
   622   \item @{text option} defines functions that map into the @{type
   623   option} type. Here, the value @{term None} is used to model a
   624   non-terminating computation. Monotonicity requires that if @{term
   625   None} is returned by a recursive call, then the overall result
   626   must also be @{term None}. This is best achieved through the use of
   627   the monadic operator @{const "Option.bind"}.
   628 
   629   \item @{text tailrec} defines functions with an arbitrary result
   630   type and uses the slightly degenerated partial order where @{term
   631   "undefined"} is the bottom element.  Now, monotonicity requires that
   632   if @{term undefined} is returned by a recursive call, then the
   633   overall result must also be @{term undefined}. In practice, this is
   634   only satisfied when each recursive call is a tail call, whose result
   635   is directly returned. Thus, this mode of operation allows the
   636   definition of arbitrary tail-recursive functions.
   637   \end{description}
   638 
   639   Experienced users may define new modes by instantiating the locale
   640   @{const "partial_function_definitions"} appropriately.
   641 
   642   \item @{attribute (HOL) partial_function_mono} declares rules for
   643   use in the internal monononicity proofs of partial function
   644   definitions.
   645 
   646   \end{description}
   647 
   648 *}
   649 
   650 subsection {* Old-style recursive function definitions (TFL) *}
   651 
   652 text {*
   653   The old TFL commands @{command (HOL) "recdef"} and @{command (HOL)
   654   "recdef_tc"} for defining recursive are mostly obsolete; @{command
   655   (HOL) "function"} or @{command (HOL) "fun"} should be used instead.
   656 
   657   \begin{matharray}{rcl}
   658     @{command_def (HOL) "recdef"} & : & @{text "theory \<rightarrow> theory)"} \\
   659     @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   660   \end{matharray}
   661 
   662   @{rail "
   663     @@{command (HOL) recdef} ('(' @'permissive' ')')? \\
   664       @{syntax name} @{syntax term} (@{syntax prop} +) hints?
   665     ;
   666     recdeftc @{syntax thmdecl}? tc
   667     ;
   668     hints: '(' @'hints' ( recdefmod * ) ')'
   669     ;
   670     recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf')
   671       (() | 'add' | 'del') ':' @{syntax thmrefs}) | @{syntax clasimpmod}
   672     ;
   673     tc: @{syntax nameref} ('(' @{syntax nat} ')')?
   674   "}
   675 
   676   \begin{description}
   677 
   678   \item @{command (HOL) "recdef"} defines general well-founded
   679   recursive functions (using the TFL package), see also
   680   \cite{isabelle-HOL}.  The ``@{text "(permissive)"}'' option tells
   681   TFL to recover from failed proof attempts, returning unfinished
   682   results.  The @{text recdef_simp}, @{text recdef_cong}, and @{text
   683   recdef_wf} hints refer to auxiliary rules to be used in the internal
   684   automated proof process of TFL.  Additional @{syntax clasimpmod}
   685   declarations (cf.\ \secref{sec:clasimp}) may be given to tune the
   686   context of the Simplifier (cf.\ \secref{sec:simplifier}) and
   687   Classical reasoner (cf.\ \secref{sec:classical}).
   688 
   689   \item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the
   690   proof for leftover termination condition number @{text i} (default
   691   1) as generated by a @{command (HOL) "recdef"} definition of
   692   constant @{text c}.
   693 
   694   Note that in most cases, @{command (HOL) "recdef"} is able to finish
   695   its internal proofs without manual intervention.
   696 
   697   \end{description}
   698 
   699   \medskip Hints for @{command (HOL) "recdef"} may be also declared
   700   globally, using the following attributes.
   701 
   702   \begin{matharray}{rcl}
   703     @{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\
   704     @{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\
   705     @{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\
   706   \end{matharray}
   707 
   708   @{rail "
   709     (@@{attribute (HOL) recdef_simp} | @@{attribute (HOL) recdef_cong} |
   710       @@{attribute (HOL) recdef_wf}) (() | 'add' | 'del')
   711   "}
   712 *}
   713 
   714 
   715 section {* Inductive and coinductive definitions \label{sec:hol-inductive} *}
   716 
   717 text {*
   718   An \textbf{inductive definition} specifies the least predicate (or
   719   set) @{text R} closed under given rules: applying a rule to elements
   720   of @{text R} yields a result within @{text R}.  For example, a
   721   structural operational semantics is an inductive definition of an
   722   evaluation relation.
   723 
   724   Dually, a \textbf{coinductive definition} specifies the greatest
   725   predicate~/ set @{text R} that is consistent with given rules: every
   726   element of @{text R} can be seen as arising by applying a rule to
   727   elements of @{text R}.  An important example is using bisimulation
   728   relations to formalise equivalence of processes and infinite data
   729   structures.
   730 
   731   \medskip The HOL package is related to the ZF one, which is
   732   described in a separate paper,\footnote{It appeared in CADE
   733   \cite{paulson-CADE}; a longer version is distributed with Isabelle.}
   734   which you should refer to in case of difficulties.  The package is
   735   simpler than that of ZF thanks to implicit type-checking in HOL.
   736   The types of the (co)inductive predicates (or sets) determine the
   737   domain of the fixedpoint definition, and the package does not have
   738   to use inference rules for type-checking.
   739 
   740   \begin{matharray}{rcl}
   741     @{command_def (HOL) "inductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   742     @{command_def (HOL) "inductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   743     @{command_def (HOL) "coinductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   744     @{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   745     @{attribute_def (HOL) mono} & : & @{text attribute} \\
   746   \end{matharray}
   747 
   748   @{rail "
   749     (@@{command (HOL) inductive} | @@{command (HOL) inductive_set} |
   750       @@{command (HOL) coinductive} | @@{command (HOL) coinductive_set})
   751     @{syntax target}? @{syntax \"fixes\"} (@'for' @{syntax \"fixes\"})? \\
   752     (@'where' clauses)? (@'monos' @{syntax thmrefs})?
   753     ;
   754     clauses: (@{syntax thmdecl}? @{syntax prop} + '|')
   755     ;
   756     @@{attribute (HOL) mono} (() | 'add' | 'del')
   757   "}
   758 
   759   \begin{description}
   760 
   761   \item @{command (HOL) "inductive"} and @{command (HOL)
   762   "coinductive"} define (co)inductive predicates from the
   763   introduction rules given in the @{keyword "where"} part.  The
   764   optional @{keyword "for"} part contains a list of parameters of the
   765   (co)inductive predicates that remain fixed throughout the
   766   definition.  The optional @{keyword "monos"} section contains
   767   \emph{monotonicity theorems}, which are required for each operator
   768   applied to a recursive set in the introduction rules.  There
   769   \emph{must} be a theorem of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"},
   770   for each premise @{text "M R\<^sub>i t"} in an introduction rule!
   771 
   772   \item @{command (HOL) "inductive_set"} and @{command (HOL)
   773   "coinductive_set"} are wrappers for to the previous commands,
   774   allowing the definition of (co)inductive sets.
   775 
   776   \item @{attribute (HOL) mono} declares monotonicity rules.  These
   777   rule are involved in the automated monotonicity proof of @{command
   778   (HOL) "inductive"}.
   779 
   780   \end{description}
   781 *}
   782 
   783 
   784 subsection {* Derived rules *}
   785 
   786 text {*
   787   Each (co)inductive definition @{text R} adds definitions to the
   788   theory and also proves some theorems:
   789 
   790   \begin{description}
   791 
   792   \item @{text R.intros} is the list of introduction rules as proven
   793   theorems, for the recursive predicates (or sets).  The rules are
   794   also available individually, using the names given them in the
   795   theory file;
   796 
   797   \item @{text R.cases} is the case analysis (or elimination) rule;
   798 
   799   \item @{text R.induct} or @{text R.coinduct} is the (co)induction
   800   rule.
   801 
   802   \end{description}
   803 
   804   When several predicates @{text "R\<^sub>1, \<dots>, R\<^sub>n"} are
   805   defined simultaneously, the list of introduction rules is called
   806   @{text "R\<^sub>1_\<dots>_R\<^sub>n.intros"}, the case analysis rules are
   807   called @{text "R\<^sub>1.cases, \<dots>, R\<^sub>n.cases"}, and the list
   808   of mutual induction rules is called @{text
   809   "R\<^sub>1_\<dots>_R\<^sub>n.inducts"}.
   810 *}
   811 
   812 
   813 subsection {* Monotonicity theorems *}
   814 
   815 text {*
   816   Each theory contains a default set of theorems that are used in
   817   monotonicity proofs.  New rules can be added to this set via the
   818   @{attribute (HOL) mono} attribute.  The HOL theory @{text Inductive}
   819   shows how this is done.  In general, the following monotonicity
   820   theorems may be added:
   821 
   822   \begin{itemize}
   823 
   824   \item Theorems of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"}, for proving
   825   monotonicity of inductive definitions whose introduction rules have
   826   premises involving terms such as @{text "M R\<^sub>i t"}.
   827 
   828   \item Monotonicity theorems for logical operators, which are of the
   829   general form @{text "(\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> (\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> \<longrightarrow> \<dots>"}.  For example, in
   830   the case of the operator @{text "\<or>"}, the corresponding theorem is
   831   \[
   832   \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"}}
   833   \]
   834 
   835   \item De Morgan style equations for reasoning about the ``polarity''
   836   of expressions, e.g.
   837   \[
   838   @{prop "\<not> \<not> P \<longleftrightarrow> P"} \qquad\qquad
   839   @{prop "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"}
   840   \]
   841 
   842   \item Equations for reducing complex operators to more primitive
   843   ones whose monotonicity can easily be proved, e.g.
   844   \[
   845   @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> \<not> P \<or> Q"} \qquad\qquad
   846   @{prop "Ball A P \<equiv> \<forall>x. x \<in> A \<longrightarrow> P x"}
   847   \]
   848 
   849   \end{itemize}
   850 
   851   %FIXME: Example of an inductive definition
   852 *}
   853 
   854 
   855 section {* Arithmetic proof support *}
   856 
   857 text {*
   858   \begin{matharray}{rcl}
   859     @{method_def (HOL) arith} & : & @{text method} \\
   860     @{attribute_def (HOL) arith} & : & @{text attribute} \\
   861     @{attribute_def (HOL) arith_split} & : & @{text attribute} \\
   862   \end{matharray}
   863 
   864   The @{method (HOL) arith} method decides linear arithmetic problems
   865   (on types @{text nat}, @{text int}, @{text real}).  Any current
   866   facts are inserted into the goal before running the procedure.
   867 
   868   The @{attribute (HOL) arith} attribute declares facts that are
   869   always supplied to the arithmetic provers implicitly.
   870 
   871   The @{attribute (HOL) arith_split} attribute declares case split
   872   rules to be expanded before @{method (HOL) arith} is invoked.
   873 
   874   Note that a simpler (but faster) arithmetic prover is
   875   already invoked by the Simplifier.
   876 *}
   877 
   878 
   879 section {* Intuitionistic proof search *}
   880 
   881 text {*
   882   \begin{matharray}{rcl}
   883     @{method_def (HOL) iprover} & : & @{text method} \\
   884   \end{matharray}
   885 
   886   @{rail "
   887     @@{method (HOL) iprover} ( @{syntax rulemod} * )
   888   "}
   889 
   890   The @{method (HOL) iprover} method performs intuitionistic proof
   891   search, depending on specifically declared rules from the context,
   892   or given as explicit arguments.  Chained facts are inserted into the
   893   goal before commencing proof search.
   894 
   895   Rules need to be classified as @{attribute (Pure) intro},
   896   @{attribute (Pure) elim}, or @{attribute (Pure) dest}; here the
   897   ``@{text "!"}'' indicator refers to ``safe'' rules, which may be
   898   applied aggressively (without considering back-tracking later).
   899   Rules declared with ``@{text "?"}'' are ignored in proof search (the
   900   single-step @{method (Pure) rule} method still observes these).  An
   901   explicit weight annotation may be given as well; otherwise the
   902   number of rule premises will be taken into account here.
   903 *}
   904 
   905 
   906 section {* Coherent Logic *}
   907 
   908 text {*
   909   \begin{matharray}{rcl}
   910     @{method_def (HOL) "coherent"} & : & @{text method} \\
   911   \end{matharray}
   912 
   913   @{rail "
   914     @@{method (HOL) coherent} @{syntax thmrefs}?
   915   "}
   916 
   917   The @{method (HOL) coherent} method solves problems of
   918   \emph{Coherent Logic} \cite{Bezem-Coquand:2005}, which covers
   919   applications in confluence theory, lattice theory and projective
   920   geometry.  See @{file "~~/src/HOL/ex/Coherent.thy"} for some
   921   examples.
   922 *}
   923 
   924 
   925 section {* Proving propositions *}
   926 
   927 text {*
   928   In addition to the standard proof methods, a number of diagnosis
   929   tools search for proofs and provide an Isar proof snippet on success.
   930   These tools are available via the following commands.
   931 
   932   \begin{matharray}{rcl}
   933     @{command_def (HOL) "solve_direct"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
   934     @{command_def (HOL) "try_methods"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
   935     @{command_def (HOL) "sledgehammer"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
   936     @{command_def (HOL) "sledgehammer_params"} & : & @{text "theory \<rightarrow> theory"}
   937   \end{matharray}
   938 
   939   @{rail "
   940     @@{command (HOL) try_methods} ( ( ( 'simp' | 'intro' | 'elim' | 'dest' ) ':' @{syntax thmrefs} ) + ) ?
   941       @{syntax nat}?
   942     ;
   943     @@{command (HOL) sledgehammer} ( '[' args ']' )? facts? @{syntax nat}?
   944     ;
   945 
   946     @@{command (HOL) sledgehammer_params} ( ( '[' args ']' ) ? )
   947     ;
   948 
   949     args: ( @{syntax name} '=' value + ',' )
   950     ;
   951 
   952     facts: '(' ( ( ( ( 'add' | 'del' ) ':' ) ? @{syntax thmrefs} ) + ) ? ')'
   953     ;
   954   "} % FIXME check args "value"
   955 
   956   \begin{description}
   957 
   958   \item @{command (HOL) "solve_direct"} checks whether the current subgoals can
   959     be solved directly by an existing theorem. Duplicate lemmas can be detected
   960     in this way.
   961 
   962   \item @{command (HOL) "try_methods"} attempts to prove a subgoal using a combination
   963     of standard proof methods (@{text auto}, @{text simp}, @{text blast}, etc.).
   964     Additional facts supplied via @{text "simp:"}, @{text "intro:"},
   965     @{text "elim:"}, and @{text "dest:"} are passed to the appropriate proof
   966     methods.
   967 
   968   \item @{command (HOL) "sledgehammer"} attempts to prove a subgoal using external
   969     automatic provers (resolution provers and SMT solvers). See the Sledgehammer
   970     manual \cite{isabelle-sledgehammer} for details.
   971 
   972   \item @{command (HOL) "sledgehammer_params"} changes
   973     @{command (HOL) "sledgehammer"} configuration options persistently.
   974 
   975   \end{description}
   976 *}
   977 
   978 
   979 section {* Checking and refuting propositions *}
   980 
   981 text {*
   982   Identifying incorrect propositions usually involves evaluation of
   983   particular assignments and systematic counterexample search.  This
   984   is supported by the following commands.
   985 
   986   \begin{matharray}{rcl}
   987     @{command_def (HOL) "value"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   988     @{command_def (HOL) "quickcheck"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
   989     @{command_def (HOL) "refute"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
   990     @{command_def (HOL) "nitpick"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
   991     @{command_def (HOL) "quickcheck_params"} & : & @{text "theory \<rightarrow> theory"} \\
   992     @{command_def (HOL) "refute_params"} & : & @{text "theory \<rightarrow> theory"} \\
   993     @{command_def (HOL) "nitpick_params"} & : & @{text "theory \<rightarrow> theory"}
   994   \end{matharray}
   995 
   996   @{rail "
   997     @@{command (HOL) value} ( '[' name ']' )? modes? @{syntax term}
   998     ;
   999 
  1000     (@@{command (HOL) quickcheck} | @@{command (HOL) refute} | @@{command (HOL) nitpick})
  1001       ( '[' args ']' )? @{syntax nat}?
  1002     ;
  1003 
  1004     (@@{command (HOL) quickcheck_params} | @@{command (HOL) refute_params} |
  1005       @@{command (HOL) nitpick_params}) ( '[' args ']' )?
  1006     ;
  1007 
  1008     modes: '(' (@{syntax name} +) ')'
  1009     ;
  1010 
  1011     args: ( @{syntax name} '=' value + ',' )
  1012     ;
  1013   "} % FIXME check "value"
  1014 
  1015   \begin{description}
  1016 
  1017   \item @{command (HOL) "value"}~@{text t} evaluates and prints a
  1018     term; optionally @{text modes} can be specified, which are
  1019     appended to the current print mode (see also \cite{isabelle-ref}).
  1020     Internally, the evaluation is performed by registered evaluators,
  1021     which are invoked sequentially until a result is returned.
  1022     Alternatively a specific evaluator can be selected using square
  1023     brackets; typical evaluators use the current set of code equations
  1024     to normalize and include @{text simp} for fully symbolic evaluation
  1025     using the simplifier, @{text nbe} for \emph{normalization by evaluation}
  1026     and \emph{code} for code generation in SML.
  1027 
  1028   \item @{command (HOL) "quickcheck"} tests the current goal for
  1029     counterexamples using a series of assignments for its
  1030     free variables; by default the first subgoal is tested, an other
  1031     can be selected explicitly using an optional goal index.
  1032     Assignments can be chosen exhausting the search space upto a given
  1033     size or using a fixed number of random assignments in the search space.
  1034     By default, quickcheck uses exhaustive testing.
  1035     A number of configuration options are supported for
  1036     @{command (HOL) "quickcheck"}, notably:
  1037 
  1038     \begin{description}
  1039 
  1040     \item[@{text tester}] specifies how to explore the search space
  1041       (e.g. exhaustive or random).
  1042       An unknown configuration option is treated as an argument to tester,
  1043       making @{text "tester ="} optional.
  1044     \item[@{text size}] specifies the maximum size of the search space
  1045     for assignment values.
  1046 
  1047     \item[@{text eval}] takes a term or a list of terms and evaluates
  1048       these terms under the variable assignment found by quickcheck.
  1049 
  1050     \item[@{text iterations}] sets how many sets of assignments are
  1051     generated for each particular size.
  1052 
  1053     \item[@{text no_assms}] specifies whether assumptions in
  1054     structured proofs should be ignored.
  1055 
  1056     \item[@{text timeout}] sets the time limit in seconds.
  1057 
  1058     \item[@{text default_type}] sets the type(s) generally used to
  1059     instantiate type variables.
  1060 
  1061     \item[@{text report}] if set quickcheck reports how many tests
  1062     fulfilled the preconditions.
  1063 
  1064     \item[@{text quiet}] if not set quickcheck informs about the
  1065     current size for assignment values.
  1066 
  1067     \item[@{text expect}] can be used to check if the user's
  1068     expectation was met (@{text no_expectation}, @{text
  1069     no_counterexample}, or @{text counterexample}).
  1070 
  1071     \end{description}
  1072 
  1073     These option can be given within square brackets.
  1074 
  1075   \item @{command (HOL) "quickcheck_params"} changes
  1076     @{command (HOL) "quickcheck"} configuration options persistently.
  1077 
  1078   \item @{command (HOL) "refute"} tests the current goal for
  1079     counterexamples using a reduction to SAT. The following configuration
  1080     options are supported:
  1081 
  1082     \begin{description}
  1083 
  1084     \item[@{text minsize}] specifies the minimum size (cardinality) of the
  1085       models to search for.
  1086 
  1087     \item[@{text maxsize}] specifies the maximum size (cardinality) of the
  1088       models to search for. Nonpositive values mean $\infty$.
  1089 
  1090     \item[@{text maxvars}] specifies the maximum number of Boolean variables
  1091     to use when transforming the term into a propositional formula.
  1092     Nonpositive values mean $\infty$.
  1093 
  1094     \item[@{text satsolver}] specifies the SAT solver to use.
  1095 
  1096     \item[@{text no_assms}] specifies whether assumptions in
  1097     structured proofs should be ignored.
  1098 
  1099     \item[@{text maxtime}] sets the time limit in seconds.
  1100 
  1101     \item[@{text expect}] can be used to check if the user's
  1102     expectation was met (@{text genuine}, @{text potential},
  1103     @{text none}, or @{text unknown}).
  1104 
  1105     \end{description}
  1106 
  1107     These option can be given within square brackets.
  1108 
  1109   \item @{command (HOL) "refute_params"} changes
  1110     @{command (HOL) "refute"} configuration options persistently.
  1111 
  1112   \item @{command (HOL) "nitpick"} tests the current goal for counterexamples
  1113     using a reduction to first-order relational logic. See the Nitpick manual
  1114     \cite{isabelle-nitpick} for details.
  1115 
  1116   \item @{command (HOL) "nitpick_params"} changes
  1117     @{command (HOL) "nitpick"} configuration options persistently.
  1118 
  1119   \end{description}
  1120 *}
  1121 
  1122 
  1123 section {* Unstructured case analysis and induction \label{sec:hol-induct-tac} *}
  1124 
  1125 text {*
  1126   The following tools of Isabelle/HOL support cases analysis and
  1127   induction in unstructured tactic scripts; see also
  1128   \secref{sec:cases-induct} for proper Isar versions of similar ideas.
  1129 
  1130   \begin{matharray}{rcl}
  1131     @{method_def (HOL) case_tac}@{text "\<^sup>*"} & : & @{text method} \\
  1132     @{method_def (HOL) induct_tac}@{text "\<^sup>*"} & : & @{text method} \\
  1133     @{method_def (HOL) ind_cases}@{text "\<^sup>*"} & : & @{text method} \\
  1134     @{command_def (HOL) "inductive_cases"}@{text "\<^sup>*"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
  1135   \end{matharray}
  1136 
  1137   @{rail "
  1138     @@{method (HOL) case_tac} @{syntax goal_spec}? @{syntax term} rule?
  1139     ;
  1140     @@{method (HOL) induct_tac} @{syntax goal_spec}? (@{syntax insts} * @'and') rule?
  1141     ;
  1142     @@{method (HOL) ind_cases} (@{syntax prop}+) (@'for' (@{syntax name}+))?
  1143     ;
  1144     @@{command (HOL) inductive_cases} (@{syntax thmdecl}? (@{syntax prop}+) + @'and')
  1145     ;
  1146 
  1147     rule: 'rule' ':' @{syntax thmref}
  1148   "}
  1149 
  1150   \begin{description}
  1151 
  1152   \item @{method (HOL) case_tac} and @{method (HOL) induct_tac} admit
  1153   to reason about inductive types.  Rules are selected according to
  1154   the declarations by the @{attribute cases} and @{attribute induct}
  1155   attributes, cf.\ \secref{sec:cases-induct}.  The @{command (HOL)
  1156   datatype} package already takes care of this.
  1157 
  1158   These unstructured tactics feature both goal addressing and dynamic
  1159   instantiation.  Note that named rule cases are \emph{not} provided
  1160   as would be by the proper @{method cases} and @{method induct} proof
  1161   methods (see \secref{sec:cases-induct}).  Unlike the @{method
  1162   induct} method, @{method induct_tac} does not handle structured rule
  1163   statements, only the compact object-logic conclusion of the subgoal
  1164   being addressed.
  1165 
  1166   \item @{method (HOL) ind_cases} and @{command (HOL)
  1167   "inductive_cases"} provide an interface to the internal @{ML_text
  1168   mk_cases} operation.  Rules are simplified in an unrestricted
  1169   forward manner.
  1170 
  1171   While @{method (HOL) ind_cases} is a proof method to apply the
  1172   result immediately as elimination rules, @{command (HOL)
  1173   "inductive_cases"} provides case split theorems at the theory level
  1174   for later use.  The @{keyword "for"} argument of the @{method (HOL)
  1175   ind_cases} method allows to specify a list of variables that should
  1176   be generalized before applying the resulting rule.
  1177 
  1178   \end{description}
  1179 *}
  1180 
  1181 
  1182 section {* Executable code *}
  1183 
  1184 text {* For validation purposes, it is often useful to \emph{execute}
  1185   specifications.  In principle, execution could be simulated by
  1186   Isabelle's inference kernel, i.e. by a combination of resolution and
  1187   simplification.  Unfortunately, this approach is rather inefficient.
  1188   A more efficient way of executing specifications is to translate
  1189   them into a functional programming language such as ML.
  1190 
  1191   Isabelle provides two generic frameworks to support code generation
  1192   from executable specifications.  Isabelle/HOL instantiates these
  1193   mechanisms in a way that is amenable to end-user applications.
  1194 *}
  1195 
  1196 
  1197 subsection {* The new code generator (F. Haftmann) *}
  1198 
  1199 text {* This framework generates code from functional programs
  1200   (including overloading using type classes) to SML \cite{SML}, OCaml
  1201   \cite{OCaml}, Haskell \cite{haskell-revised-report} and Scala
  1202   \cite{scala-overview-tech-report}.  Conceptually, code generation is
  1203   split up in three steps: \emph{selection} of code theorems,
  1204   \emph{translation} into an abstract executable view and
  1205   \emph{serialization} to a specific \emph{target language}.
  1206   Inductive specifications can be executed using the predicate
  1207   compiler which operates within HOL.  See \cite{isabelle-codegen} for
  1208   an introduction.
  1209 
  1210   \begin{matharray}{rcl}
  1211     @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1212     @{attribute_def (HOL) code} & : & @{text attribute} \\
  1213     @{command_def (HOL) "code_abort"} & : & @{text "theory \<rightarrow> theory"} \\
  1214     @{command_def (HOL) "code_datatype"} & : & @{text "theory \<rightarrow> theory"} \\
  1215     @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1216     @{attribute_def (HOL) code_inline} & : & @{text attribute} \\
  1217     @{attribute_def (HOL) code_post} & : & @{text attribute} \\
  1218     @{command_def (HOL) "print_codeproc"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1219     @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1220     @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1221     @{command_def (HOL) "code_const"} & : & @{text "theory \<rightarrow> theory"} \\
  1222     @{command_def (HOL) "code_type"} & : & @{text "theory \<rightarrow> theory"} \\
  1223     @{command_def (HOL) "code_class"} & : & @{text "theory \<rightarrow> theory"} \\
  1224     @{command_def (HOL) "code_instance"} & : & @{text "theory \<rightarrow> theory"} \\
  1225     @{command_def (HOL) "code_reserved"} & : & @{text "theory \<rightarrow> theory"} \\
  1226     @{command_def (HOL) "code_monad"} & : & @{text "theory \<rightarrow> theory"} \\
  1227     @{command_def (HOL) "code_include"} & : & @{text "theory \<rightarrow> theory"} \\
  1228     @{command_def (HOL) "code_modulename"} & : & @{text "theory \<rightarrow> theory"} \\
  1229     @{command_def (HOL) "code_reflect"} & : & @{text "theory \<rightarrow> theory"}
  1230   \end{matharray}
  1231 
  1232   @{rail "
  1233     @@{command (HOL) export_code} ( constexpr + ) \\
  1234        ( ( @'in' target ( @'module_name' @{syntax string} ) ? \\
  1235         ( @'file' ( @{syntax string} | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
  1236     ;
  1237 
  1238     const: @{syntax term}
  1239     ;
  1240 
  1241     constexpr: ( const | 'name._' | '_' )
  1242     ;
  1243 
  1244     typeconstructor: @{syntax nameref}
  1245     ;
  1246 
  1247     class: @{syntax nameref}
  1248     ;
  1249 
  1250     target: 'SML' | 'OCaml' | 'Haskell' | 'Scala'
  1251     ;
  1252 
  1253     @@{attribute (HOL) code} ( 'del' | 'abstype' | 'abstract' )?
  1254     ;
  1255 
  1256     @@{command (HOL) code_abort} ( const + )
  1257     ;
  1258 
  1259     @@{command (HOL) code_datatype} ( const + )
  1260     ;
  1261 
  1262     @@{attribute (HOL) code_inline} ( 'del' ) ?
  1263     ;
  1264 
  1265     @@{attribute (HOL) code_post} ( 'del' ) ?
  1266     ;
  1267 
  1268     @@{command (HOL) code_thms} ( constexpr + ) ?
  1269     ;
  1270 
  1271     @@{command (HOL) code_deps} ( constexpr + ) ?
  1272     ;
  1273 
  1274     @@{command (HOL) code_const} (const + @'and') \\
  1275       ( ( '(' target ( syntax ? + @'and' ) ')' ) + )
  1276     ;
  1277 
  1278     @@{command (HOL) code_type} (typeconstructor + @'and') \\
  1279       ( ( '(' target ( syntax ? + @'and' ) ')' ) + )
  1280     ;
  1281 
  1282     @@{command (HOL) code_class} (class + @'and') \\
  1283       ( ( '(' target \\ ( @{syntax string} ? + @'and' ) ')' ) + )
  1284     ;
  1285 
  1286     @@{command (HOL) code_instance} (( typeconstructor '::' class ) + @'and') \\
  1287       ( ( '(' target ( '-' ? + @'and' ) ')' ) + )
  1288     ;
  1289 
  1290     @@{command (HOL) code_reserved} target ( @{syntax string} + )
  1291     ;
  1292 
  1293     @@{command (HOL) code_monad} const const target
  1294     ;
  1295 
  1296     @@{command (HOL) code_include} target ( @{syntax string} ( @{syntax string} | '-') )
  1297     ;
  1298 
  1299     @@{command (HOL) code_modulename} target ( ( @{syntax string} @{syntax string} ) + )
  1300     ;
  1301 
  1302     @@{command (HOL) code_reflect} @{syntax string} \\
  1303       ( @'datatypes' ( @{syntax string} '=' ( '_' | ( @{syntax string} + '|' ) + @'and' ) ) ) ? \\
  1304       ( @'functions' ( @{syntax string} + ) ) ? ( @'file' @{syntax string} ) ?
  1305     ;
  1306 
  1307     syntax: @{syntax string} | ( @'infix' | @'infixl' | @'infixr' ) @{syntax nat} @{syntax string}
  1308   "}
  1309 
  1310   \begin{description}
  1311 
  1312   \item @{command (HOL) "export_code"} generates code for a given list
  1313   of constants in the specified target language(s).  If no
  1314   serialization instruction is given, only abstract code is generated
  1315   internally.
  1316 
  1317   Constants may be specified by giving them literally, referring to
  1318   all executable contants within a certain theory by giving @{text
  1319   "name.*"}, or referring to \emph{all} executable constants currently
  1320   available by giving @{text "*"}.
  1321 
  1322   By default, for each involved theory one corresponding name space
  1323   module is generated.  Alternativly, a module name may be specified
  1324   after the @{keyword "module_name"} keyword; then \emph{all} code is
  1325   placed in this module.
  1326 
  1327   For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification
  1328   refers to a single file; for \emph{Haskell}, it refers to a whole
  1329   directory, where code is generated in multiple files reflecting the
  1330   module hierarchy.  Omitting the file specification denotes standard
  1331   output.
  1332 
  1333   Serializers take an optional list of arguments in parentheses.  For
  1334   \emph{SML} and \emph{OCaml}, ``@{text no_signatures}`` omits
  1335   explicit module signatures.
  1336 
  1337   For \emph{Haskell} a module name prefix may be given using the
  1338   ``@{text "root:"}'' argument; ``@{text string_classes}'' adds a
  1339   ``@{verbatim "deriving (Read, Show)"}'' clause to each appropriate
  1340   datatype declaration.
  1341 
  1342   \item @{attribute (HOL) code} explicitly selects (or with option
  1343   ``@{text "del"}'' deselects) a code equation for code generation.
  1344   Usually packages introducing code equations provide a reasonable
  1345   default setup for selection.  Variants @{text "code abstype"} and
  1346   @{text "code abstract"} declare abstract datatype certificates or
  1347   code equations on abstract datatype representations respectively.
  1348 
  1349   \item @{command (HOL) "code_abort"} declares constants which are not
  1350   required to have a definition by means of code equations; if needed
  1351   these are implemented by program abort instead.
  1352 
  1353   \item @{command (HOL) "code_datatype"} specifies a constructor set
  1354   for a logical type.
  1355 
  1356   \item @{command (HOL) "print_codesetup"} gives an overview on
  1357   selected code equations and code generator datatypes.
  1358 
  1359   \item @{attribute (HOL) code_inline} declares (or with option
  1360   ``@{text "del"}'' removes) inlining theorems which are applied as
  1361   rewrite rules to any code equation during preprocessing.
  1362 
  1363   \item @{attribute (HOL) code_post} declares (or with option ``@{text
  1364   "del"}'' removes) theorems which are applied as rewrite rules to any
  1365   result of an evaluation.
  1366 
  1367   \item @{command (HOL) "print_codeproc"} prints the setup of the code
  1368   generator preprocessor.
  1369 
  1370   \item @{command (HOL) "code_thms"} prints a list of theorems
  1371   representing the corresponding program containing all given
  1372   constants after preprocessing.
  1373 
  1374   \item @{command (HOL) "code_deps"} visualizes dependencies of
  1375   theorems representing the corresponding program containing all given
  1376   constants after preprocessing.
  1377 
  1378   \item @{command (HOL) "code_const"} associates a list of constants
  1379   with target-specific serializations; omitting a serialization
  1380   deletes an existing serialization.
  1381 
  1382   \item @{command (HOL) "code_type"} associates a list of type
  1383   constructors with target-specific serializations; omitting a
  1384   serialization deletes an existing serialization.
  1385 
  1386   \item @{command (HOL) "code_class"} associates a list of classes
  1387   with target-specific class names; omitting a serialization deletes
  1388   an existing serialization.  This applies only to \emph{Haskell}.
  1389 
  1390   \item @{command (HOL) "code_instance"} declares a list of type
  1391   constructor / class instance relations as ``already present'' for a
  1392   given target.  Omitting a ``@{text "-"}'' deletes an existing
  1393   ``already present'' declaration.  This applies only to
  1394   \emph{Haskell}.
  1395 
  1396   \item @{command (HOL) "code_reserved"} declares a list of names as
  1397   reserved for a given target, preventing it to be shadowed by any
  1398   generated code.
  1399 
  1400   \item @{command (HOL) "code_monad"} provides an auxiliary mechanism
  1401   to generate monadic code for Haskell.
  1402 
  1403   \item @{command (HOL) "code_include"} adds arbitrary named content
  1404   (``include'') to generated code.  A ``@{text "-"}'' as last argument
  1405   will remove an already added ``include''.
  1406 
  1407   \item @{command (HOL) "code_modulename"} declares aliasings from one
  1408   module name onto another.
  1409 
  1410   \item @{command (HOL) "code_reflect"} without a ``@{text "file"}''
  1411   argument compiles code into the system runtime environment and
  1412   modifies the code generator setup that future invocations of system
  1413   runtime code generation referring to one of the ``@{text
  1414   "datatypes"}'' or ``@{text "functions"}'' entities use these precompiled
  1415   entities.  With a ``@{text "file"}'' argument, the corresponding code
  1416   is generated into that specified file without modifying the code
  1417   generator setup.
  1418 
  1419   \end{description}
  1420 *}
  1421 
  1422 
  1423 subsection {* The old code generator (S. Berghofer) *}
  1424 
  1425 text {* This framework generates code from both functional and
  1426   relational programs to SML, as explained below.
  1427 
  1428   \begin{matharray}{rcl}
  1429     @{command_def "code_module"} & : & @{text "theory \<rightarrow> theory"} \\
  1430     @{command_def "code_library"} & : & @{text "theory \<rightarrow> theory"} \\
  1431     @{command_def "consts_code"} & : & @{text "theory \<rightarrow> theory"} \\
  1432     @{command_def "types_code"} & : & @{text "theory \<rightarrow> theory"} \\
  1433     @{attribute_def code} & : & @{text attribute} \\
  1434   \end{matharray}
  1435 
  1436   @{rail "
  1437   ( @@{command code_module} | @@{command code_library} ) modespec? @{syntax name}? \\
  1438     ( @'file' name ) ? ( @'imports' ( @{syntax name} + ) ) ? \\
  1439     @'contains' ( ( @{syntax name} '=' @{syntax term} ) + | @{syntax term} + )
  1440   ;
  1441 
  1442   modespec: '(' ( @{syntax name} * ) ')'
  1443   ;
  1444 
  1445   @@{command (HOL) consts_code} (codespec +)
  1446   ;
  1447 
  1448   codespec: const template attachment ?
  1449   ;
  1450 
  1451   @@{command (HOL) types_code} (tycodespec +)
  1452   ;
  1453 
  1454   tycodespec: @{syntax name} template attachment ?
  1455   ;
  1456 
  1457   const: @{syntax term}
  1458   ;
  1459 
  1460   template: '(' @{syntax string} ')'
  1461   ;
  1462 
  1463   attachment: 'attach' modespec? '{' @{syntax text} '}'
  1464   ;
  1465 
  1466   @@{attribute code} name?
  1467   "}
  1468 *}
  1469 
  1470 
  1471 subsubsection {* Invoking the code generator *}
  1472 
  1473 text {* The code generator is invoked via the @{command code_module}
  1474   and @{command code_library} commands, which correspond to
  1475   \emph{incremental} and \emph{modular} code generation, respectively.
  1476 
  1477   \begin{description}
  1478 
  1479   \item [Modular] For each theory, an ML structure is generated,
  1480   containing the code generated from the constants defined in this
  1481   theory.
  1482 
  1483   \item [Incremental] All the generated code is emitted into the same
  1484   structure.  This structure may import code from previously generated
  1485   structures, which can be specified via @{keyword "imports"}.
  1486   Moreover, the generated structure may also be referred to in later
  1487   invocations of the code generator.
  1488 
  1489   \end{description}
  1490 
  1491   After the @{command code_module} and @{command code_library}
  1492   keywords, the user may specify an optional list of ``modes'' in
  1493   parentheses. These can be used to instruct the code generator to
  1494   emit additional code for special purposes, e.g.\ functions for
  1495   converting elements of generated datatypes to Isabelle terms, or
  1496   test data generators. The list of modes is followed by a module
  1497   name.  The module name is optional for modular code generation, but
  1498   must be specified for incremental code generation.
  1499 
  1500   The code can either be written to a file, in which case a file name
  1501   has to be specified after the @{keyword "file"} keyword, or be loaded
  1502   directly into Isabelle's ML environment. In the latter case, the
  1503   @{command ML} theory command can be used to inspect the results
  1504   interactively, for example.
  1505 
  1506   The terms from which to generate code can be specified after the
  1507   @{keyword "contains"} keyword, either as a list of bindings, or just
  1508   as a list of terms. In the latter case, the code generator just
  1509   produces code for all constants and types occuring in the term, but
  1510   does not bind the compiled terms to ML identifiers.
  1511 
  1512   Here is an example:
  1513 *}
  1514 
  1515 code_module Test
  1516 contains test = "foldl op + (0 :: int) [1, 2, 3, 4, 5]"
  1517 
  1518 text {* \noindent This binds the result of compiling the given term to
  1519   the ML identifier @{ML Test.test}.  *}
  1520 
  1521 ML {* @{assert} (Test.test = 15) *}
  1522 
  1523 
  1524 subsubsection {* Configuring the code generator *}
  1525 
  1526 text {* When generating code for a complex term, the code generator
  1527   recursively calls itself for all subterms.  When it arrives at a
  1528   constant, the default strategy of the code generator is to look up
  1529   its definition and try to generate code for it.  Constants which
  1530   have no definitions that are immediately executable, may be
  1531   associated with a piece of ML code manually using the @{command_ref
  1532   consts_code} command.  It takes a list whose elements consist of a
  1533   constant (given in usual term syntax -- an explicit type constraint
  1534   accounts for overloading), and a mixfix template describing the ML
  1535   code. The latter is very much the same as the mixfix templates used
  1536   when declaring new constants.  The most notable difference is that
  1537   terms may be included in the ML template using antiquotation
  1538   brackets @{verbatim "{"}@{verbatim "*"}~@{text "..."}~@{verbatim
  1539   "*"}@{verbatim "}"}.
  1540 
  1541   A similar mechanism is available for types: @{command_ref
  1542   types_code} associates type constructors with specific ML code.
  1543 
  1544   For example, the following declarations copied from @{file
  1545   "~~/src/HOL/Product_Type.thy"} describe how the product type of
  1546   Isabelle/HOL should be compiled to ML.  *}
  1547 
  1548 typedecl ('a, 'b) prod
  1549 consts Pair :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) prod"
  1550 
  1551 types_code prod  ("(_ */ _)")
  1552 consts_code Pair  ("(_,/ _)")
  1553 
  1554 text {* Sometimes, the code associated with a constant or type may
  1555   need to refer to auxiliary functions, which have to be emitted when
  1556   the constant is used. Code for such auxiliary functions can be
  1557   declared using @{keyword "attach"}. For example, the @{const wfrec}
  1558   function can be implemented as follows:
  1559 *}
  1560 
  1561 consts_code wfrec  ("\<module>wfrec?")  (* FIXME !? *)
  1562 attach {* fun wfrec f x = f (wfrec f) x *}
  1563 
  1564 text {* If the code containing a call to @{const wfrec} resides in an
  1565   ML structure different from the one containing the function
  1566   definition attached to @{const wfrec}, the name of the ML structure
  1567   (followed by a ``@{text "."}'')  is inserted in place of ``@{text
  1568   "\<module>"}'' in the above template.  The ``@{text "?"}''  means that
  1569   the code generator should ignore the first argument of @{const
  1570   wfrec}, i.e.\ the termination relation, which is usually not
  1571   executable.
  1572 
  1573   \medskip Another possibility of configuring the code generator is to
  1574   register theorems to be used for code generation. Theorems can be
  1575   registered via the @{attribute code} attribute. It takes an optional
  1576   name as an argument, which indicates the format of the
  1577   theorem. Currently supported formats are equations (this is the
  1578   default when no name is specified) and horn clauses (this is
  1579   indicated by the name \texttt{ind}). The left-hand sides of
  1580   equations may only contain constructors and distinct variables,
  1581   whereas horn clauses must have the same format as introduction rules
  1582   of inductive definitions.
  1583 
  1584   The following example specifies three equations from which to
  1585   generate code for @{term "op <"} on natural numbers (see also
  1586   @{"file" "~~/src/HOL/Nat.thy"}).  *}
  1587 
  1588 lemma [code]: "(Suc m < Suc n) = (m < n)"
  1589   and [code]: "((n::nat) < 0) = False"
  1590   and [code]: "(0 < Suc n) = True" by simp_all
  1591 
  1592 
  1593 subsubsection {* Specific HOL code generators *}
  1594 
  1595 text {* The basic code generator framework offered by Isabelle/Pure
  1596   has already been extended with additional code generators for
  1597   specific HOL constructs. These include datatypes, recursive
  1598   functions and inductive relations. The code generator for inductive
  1599   relations can handle expressions of the form @{text "(t\<^sub>1, \<dots>, t\<^sub>n) \<in>
  1600   r"}, where @{text "r"} is an inductively defined relation. If at
  1601   least one of the @{text "t\<^sub>i"} is a dummy pattern ``@{text "_"}'',
  1602   the above expression evaluates to a sequence of possible answers. If
  1603   all of the @{text "t\<^sub>i"} are proper terms, the expression evaluates
  1604   to a boolean value.
  1605 
  1606   The following example demonstrates this for beta-reduction on lambda
  1607   terms (see also @{"file" "~~/src/HOL/Proofs/Lambda/Lambda.thy"}).
  1608 *}
  1609 
  1610 datatype dB =
  1611     Var nat
  1612   | App dB dB  (infixl "\<degree>" 200)
  1613   | Abs dB
  1614 
  1615 primrec lift :: "dB \<Rightarrow> nat \<Rightarrow> dB"
  1616 where
  1617     "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
  1618   | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
  1619   | "lift (Abs s) k = Abs (lift s (k + 1))"
  1620 
  1621 primrec subst :: "dB \<Rightarrow> dB \<Rightarrow> nat \<Rightarrow> dB"  ("_[_'/_]" [300, 0, 0] 300)
  1622 where
  1623     "(Var i)[s/k] =
  1624       (if k < i then Var (i - 1) else if i = k then s else Var i)"
  1625   | "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
  1626   | "(Abs t)[s/k] = Abs (t[lift s 0 / k+1])"
  1627 
  1628 inductive beta :: "dB \<Rightarrow> dB \<Rightarrow> bool"  (infixl "\<rightarrow>\<^sub>\<beta>" 50)
  1629 where
  1630     beta: "Abs s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
  1631   | appL: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
  1632   | appR: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
  1633   | abs: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> Abs s \<rightarrow>\<^sub>\<beta> Abs t"
  1634 
  1635 code_module Test
  1636 contains
  1637   test1 = "Abs (Var 0) \<degree> Var 0 \<rightarrow>\<^sub>\<beta> Var 0"
  1638   test2 = "Abs (Abs (Var 0 \<degree> Var 0) \<degree> (Abs (Var 0) \<degree> Var 0)) \<rightarrow>\<^sub>\<beta> _"
  1639 
  1640 text {*
  1641   In the above example, @{ML Test.test1} evaluates to a boolean,
  1642   whereas @{ML Test.test2} is a lazy sequence whose elements can be
  1643   inspected separately.
  1644 *}
  1645 
  1646 ML {* @{assert} Test.test1 *}
  1647 ML {* val results = DSeq.list_of Test.test2 *}
  1648 ML {* @{assert} (length results = 2) *}
  1649 
  1650 text {*
  1651   \medskip The theory underlying the HOL code generator is described
  1652   more detailed in \cite{Berghofer-Nipkow:2002}. More examples that
  1653   illustrate the usage of the code generator can be found e.g.\ in
  1654   @{"file" "~~/src/HOL/MicroJava/J/JListExample.thy"} and @{"file"
  1655   "~~/src/HOL/MicroJava/JVM/JVMListExample.thy"}.
  1656 *}
  1657 
  1658 
  1659 section {* Definition by specification \label{sec:hol-specification} *}
  1660 
  1661 text {*
  1662   \begin{matharray}{rcl}
  1663     @{command_def (HOL) "specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
  1664     @{command_def (HOL) "ax_specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
  1665   \end{matharray}
  1666 
  1667   @{rail "
  1668   (@@{command (HOL) specification} | @@{command (HOL) ax_specification})
  1669     '(' (decl +) ')' \\ (@{syntax thmdecl}? @{syntax prop} +)
  1670   ;
  1671   decl: ((@{syntax name} ':')? @{syntax term} '(' @'overloaded' ')'?)
  1672   "}
  1673 
  1674   \begin{description}
  1675 
  1676   \item @{command (HOL) "specification"}~@{text "decls \<phi>"} sets up a
  1677   goal stating the existence of terms with the properties specified to
  1678   hold for the constants given in @{text decls}.  After finishing the
  1679   proof, the theory will be augmented with definitions for the given
  1680   constants, as well as with theorems stating the properties for these
  1681   constants.
  1682 
  1683   \item @{command (HOL) "ax_specification"}~@{text "decls \<phi>"} sets up
  1684   a goal stating the existence of terms with the properties specified
  1685   to hold for the constants given in @{text decls}.  After finishing
  1686   the proof, the theory will be augmented with axioms expressing the
  1687   properties given in the first place.
  1688 
  1689   \item @{text decl} declares a constant to be defined by the
  1690   specification given.  The definition for the constant @{text c} is
  1691   bound to the name @{text c_def} unless a theorem name is given in
  1692   the declaration.  Overloaded constants should be declared as such.
  1693 
  1694   \end{description}
  1695 
  1696   Whether to use @{command (HOL) "specification"} or @{command (HOL)
  1697   "ax_specification"} is to some extent a matter of style.  @{command
  1698   (HOL) "specification"} introduces no new axioms, and so by
  1699   construction cannot introduce inconsistencies, whereas @{command
  1700   (HOL) "ax_specification"} does introduce axioms, but only after the
  1701   user has explicitly proven it to be safe.  A practical issue must be
  1702   considered, though: After introducing two constants with the same
  1703   properties using @{command (HOL) "specification"}, one can prove
  1704   that the two constants are, in fact, equal.  If this might be a
  1705   problem, one should use @{command (HOL) "ax_specification"}.
  1706 *}
  1707 
  1708 end