doc-src/IsarRef/Thy/HOL_Specific.thy
author wenzelm
Sat, 26 Mar 2011 16:10:22 +0100
changeset 42994 c407078c0d47
parent 42956 f07b373f25d3
child 43082 de9d43c427ae
permissions -rw-r--r--
updated generated file;
tuned whitespace;
     1 theory HOL_Specific
     2 imports 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   \begin{rail}
    15     'typedef' altname? abstype '=' repset
    16     ;
    17 
    18     altname: '(' (name | 'open' | 'open' name) ')'
    19     ;
    20     abstype: typespecsorts mixfix?
    21     ;
    22     repset: term ('morphisms' name name)?
    23     ;
    24   \end{rail}
    25 
    26   \begin{description}
    27 
    28   \item @{command (HOL) "typedef"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = A"}
    29   axiomatizes a Gordon/HOL-style type definition in the background
    30   theory of the current context, depending on a non-emptiness result
    31   of the set @{text A} (which needs to be proven interactively).
    32 
    33   The raw type may not depend on parameters or assumptions of the
    34   context --- this is logically impossible in Isabelle/HOL --- but the
    35   non-emptiness property can be local, potentially resulting in
    36   multiple interpretations in target contexts.  Thus the established
    37   bijection between the representing set @{text A} and the new type
    38   @{text t} may semantically depend on local assumptions.
    39 
    40   By default, @{command (HOL) "typedef"} defines both a type @{text t}
    41   and a set (term constant) of the same name, unless an alternative
    42   base name is given in parentheses, or the ``@{text "(open)"}''
    43   declaration is used to suppress a separate constant definition
    44   altogether.  The injection from type to set is called @{text Rep_t},
    45   its inverse @{text Abs_t} --- this may be changed via an explicit
    46   @{keyword (HOL) "morphisms"} declaration.
    47 
    48   Theorems @{text Rep_t}, @{text Rep_t_inverse}, and @{text
    49   Abs_t_inverse} provide the most basic characterization as a
    50   corresponding injection/surjection pair (in both directions).  Rules
    51   @{text Rep_t_inject} and @{text Abs_t_inject} provide a slightly
    52   more convenient view on the injectivity part, suitable for automated
    53   proof tools (e.g.\ in @{attribute simp} or @{attribute iff}
    54   declarations).  Rules @{text Rep_t_cases}/@{text Rep_t_induct}, and
    55   @{text Abs_t_cases}/@{text Abs_t_induct} provide alternative views
    56   on surjectivity; these are already declared as set or type rules for
    57   the generic @{method cases} and @{method induct} methods.
    58 
    59   An alternative name for the set definition (and other derived
    60   entities) may be specified in parentheses; the default is to use
    61   @{text t} as indicated before.
    62 
    63   \end{description}
    64 *}
    65 
    66 
    67 section {* Adhoc tuples *}
    68 
    69 text {*
    70   \begin{matharray}{rcl}
    71     @{attribute (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\
    72   \end{matharray}
    73 
    74   \begin{rail}
    75     'split_format' '(' 'complete' ')'
    76     ;
    77   \end{rail}
    78 
    79   \begin{description}
    80 
    81   \item @{attribute (HOL) split_format}\ @{text "(complete)"} causes
    82   arguments in function applications to be represented canonically
    83   according to their tuple type structure.
    84 
    85   Note that this operation tends to invent funny names for new local
    86   parameters introduced.
    87 
    88   \end{description}
    89 *}
    90 
    91 
    92 section {* Records \label{sec:hol-record} *}
    93 
    94 text {*
    95   In principle, records merely generalize the concept of tuples, where
    96   components may be addressed by labels instead of just position.  The
    97   logical infrastructure of records in Isabelle/HOL is slightly more
    98   advanced, though, supporting truly extensible record schemes.  This
    99   admits operations that are polymorphic with respect to record
   100   extension, yielding ``object-oriented'' effects like (single)
   101   inheritance.  See also \cite{NaraschewskiW-TPHOLs98} for more
   102   details on object-oriented verification and record subtyping in HOL.
   103 *}
   104 
   105 
   106 subsection {* Basic concepts *}
   107 
   108 text {*
   109   Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records
   110   at the level of terms and types.  The notation is as follows:
   111 
   112   \begin{center}
   113   \begin{tabular}{l|l|l}
   114     & record terms & record types \\ \hline
   115     fixed & @{text "\<lparr>x = a, y = b\<rparr>"} & @{text "\<lparr>x :: A, y :: B\<rparr>"} \\
   116     schematic & @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} &
   117       @{text "\<lparr>x :: A, y :: B, \<dots> :: M\<rparr>"} \\
   118   \end{tabular}
   119   \end{center}
   120 
   121   \noindent The ASCII representation of @{text "\<lparr>x = a\<rparr>"} is @{text
   122   "(| x = a |)"}.
   123 
   124   A fixed record @{text "\<lparr>x = a, y = b\<rparr>"} has field @{text x} of value
   125   @{text a} and field @{text y} of value @{text b}.  The corresponding
   126   type is @{text "\<lparr>x :: A, y :: B\<rparr>"}, assuming that @{text "a :: A"}
   127   and @{text "b :: B"}.
   128 
   129   A record scheme like @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} contains fields
   130   @{text x} and @{text y} as before, but also possibly further fields
   131   as indicated by the ``@{text "\<dots>"}'' notation (which is actually part
   132   of the syntax).  The improper field ``@{text "\<dots>"}'' of a record
   133   scheme is called the \emph{more part}.  Logically it is just a free
   134   variable, which is occasionally referred to as ``row variable'' in
   135   the literature.  The more part of a record scheme may be
   136   instantiated by zero or more further components.  For example, the
   137   previous scheme may get instantiated to @{text "\<lparr>x = a, y = b, z =
   138   c, \<dots> = m'\<rparr>"}, where @{text m'} refers to a different more part.
   139   Fixed records are special instances of record schemes, where
   140   ``@{text "\<dots>"}'' is properly terminated by the @{text "() :: unit"}
   141   element.  In fact, @{text "\<lparr>x = a, y = b\<rparr>"} is just an abbreviation
   142   for @{text "\<lparr>x = a, y = b, \<dots> = ()\<rparr>"}.
   143 
   144   \medskip Two key observations make extensible records in a simply
   145   typed language like HOL work out:
   146 
   147   \begin{enumerate}
   148 
   149   \item the more part is internalized, as a free term or type
   150   variable,
   151 
   152   \item field names are externalized, they cannot be accessed within
   153   the logic as first-class values.
   154 
   155   \end{enumerate}
   156 
   157   \medskip In Isabelle/HOL record types have to be defined explicitly,
   158   fixing their field names and types, and their (optional) parent
   159   record.  Afterwards, records may be formed using above syntax, while
   160   obeying the canonical order of fields as given by their declaration.
   161   The record package provides several standard operations like
   162   selectors and updates.  The common setup for various generic proof
   163   tools enable succinct reasoning patterns.  See also the Isabelle/HOL
   164   tutorial \cite{isabelle-hol-book} for further instructions on using
   165   records in practice.
   166 *}
   167 
   168 
   169 subsection {* Record specifications *}
   170 
   171 text {*
   172   \begin{matharray}{rcl}
   173     @{command_def (HOL) "record"} & : & @{text "theory \<rightarrow> theory"} \\
   174   \end{matharray}
   175 
   176   \begin{rail}
   177     'record' typespecsorts '=' (type '+')? (constdecl +)
   178     ;
   179   \end{rail}
   180 
   181   \begin{description}
   182 
   183   \item @{command (HOL) "record"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t = \<tau> + c\<^sub>1 :: \<sigma>\<^sub>1
   184   \<dots> c\<^sub>n :: \<sigma>\<^sub>n"} defines extensible record type @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"},
   185   derived from the optional parent record @{text "\<tau>"} by adding new
   186   field components @{text "c\<^sub>i :: \<sigma>\<^sub>i"} etc.
   187 
   188   The type variables of @{text "\<tau>"} and @{text "\<sigma>\<^sub>i"} need to be
   189   covered by the (distinct) parameters @{text "\<alpha>\<^sub>1, \<dots>,
   190   \<alpha>\<^sub>m"}.  Type constructor @{text t} has to be new, while @{text
   191   \<tau>} needs to specify an instance of an existing record type.  At
   192   least one new field @{text "c\<^sub>i"} has to be specified.
   193   Basically, field names need to belong to a unique record.  This is
   194   not a real restriction in practice, since fields are qualified by
   195   the record name internally.
   196 
   197   The parent record specification @{text \<tau>} is optional; if omitted
   198   @{text t} becomes a root record.  The hierarchy of all records
   199   declared within a theory context forms a forest structure, i.e.\ a
   200   set of trees starting with a root record each.  There is no way to
   201   merge multiple parent records!
   202 
   203   For convenience, @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} is made a
   204   type abbreviation for the fixed record type @{text "\<lparr>c\<^sub>1 ::
   205   \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n\<rparr>"}, likewise is @{text
   206   "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m, \<zeta>) t_scheme"} made an abbreviation for
   207   @{text "\<lparr>c\<^sub>1 :: \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n, \<dots> ::
   208   \<zeta>\<rparr>"}.
   209 
   210   \end{description}
   211 *}
   212 
   213 
   214 subsection {* Record operations *}
   215 
   216 text {*
   217   Any record definition of the form presented above produces certain
   218   standard operations.  Selectors and updates are provided for any
   219   field, including the improper one ``@{text more}''.  There are also
   220   cumulative record constructor functions.  To simplify the
   221   presentation below, we assume for now that @{text "(\<alpha>\<^sub>1, \<dots>,
   222   \<alpha>\<^sub>m) t"} is a root record with fields @{text "c\<^sub>1 ::
   223   \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n"}.
   224 
   225   \medskip \textbf{Selectors} and \textbf{updates} are available for
   226   any field (including ``@{text more}''):
   227 
   228   \begin{matharray}{lll}
   229     @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
   230     @{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>"} \\
   231   \end{matharray}
   232 
   233   There is special syntax for application of updates: @{text "r\<lparr>x :=
   234   a\<rparr>"} abbreviates term @{text "x_update a r"}.  Further notation for
   235   repeated updates is also available: @{text "r\<lparr>x := a\<rparr>\<lparr>y := b\<rparr>\<lparr>z :=
   236   c\<rparr>"} may be written @{text "r\<lparr>x := a, y := b, z := c\<rparr>"}.  Note that
   237   because of postfix notation the order of fields shown here is
   238   reverse than in the actual term.  Since repeated updates are just
   239   function applications, fields may be freely permuted in @{text "\<lparr>x
   240   := a, y := b, z := c\<rparr>"}, as far as logical equality is concerned.
   241   Thus commutativity of independent updates can be proven within the
   242   logic for any two fields, but not as a general theorem.
   243 
   244   \medskip The \textbf{make} operation provides a cumulative record
   245   constructor function:
   246 
   247   \begin{matharray}{lll}
   248     @{text "t.make"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   249   \end{matharray}
   250 
   251   \medskip We now reconsider the case of non-root records, which are
   252   derived of some parent.  In general, the latter may depend on
   253   another parent as well, resulting in a list of \emph{ancestor
   254   records}.  Appending the lists of fields of all ancestors results in
   255   a certain field prefix.  The record package automatically takes care
   256   of this by lifting operations over this context of ancestor fields.
   257   Assuming that @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} has ancestor
   258   fields @{text "b\<^sub>1 :: \<rho>\<^sub>1, \<dots>, b\<^sub>k :: \<rho>\<^sub>k"},
   259   the above record operations will get the following types:
   260 
   261   \medskip
   262   \begin{tabular}{lll}
   263     @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
   264     @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow>
   265       \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow>
   266       \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
   267     @{text "t.make"} & @{text "::"} & @{text "\<rho>\<^sub>1 \<Rightarrow> \<dots> \<rho>\<^sub>k \<Rightarrow> \<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow>
   268       \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   269   \end{tabular}
   270   \medskip
   271 
   272   \noindent Some further operations address the extension aspect of a
   273   derived record scheme specifically: @{text "t.fields"} produces a
   274   record fragment consisting of exactly the new fields introduced here
   275   (the result may serve as a more part elsewhere); @{text "t.extend"}
   276   takes a fixed record and adds a given more part; @{text
   277   "t.truncate"} restricts a record scheme to a fixed record.
   278 
   279   \medskip
   280   \begin{tabular}{lll}
   281     @{text "t.fields"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   282     @{text "t.extend"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr> \<Rightarrow>
   283       \<zeta> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
   284     @{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>"} \\
   285   \end{tabular}
   286   \medskip
   287 
   288   \noindent Note that @{text "t.make"} and @{text "t.fields"} coincide
   289   for root records.
   290 *}
   291 
   292 
   293 subsection {* Derived rules and proof tools *}
   294 
   295 text {*
   296   The record package proves several results internally, declaring
   297   these facts to appropriate proof tools.  This enables users to
   298   reason about record structures quite conveniently.  Assume that
   299   @{text t} is a record type as specified above.
   300 
   301   \begin{enumerate}
   302 
   303   \item Standard conversions for selectors or updates applied to
   304   record constructor terms are made part of the default Simplifier
   305   context; thus proofs by reduction of basic operations merely require
   306   the @{method simp} method without further arguments.  These rules
   307   are available as @{text "t.simps"}, too.
   308 
   309   \item Selectors applied to updated records are automatically reduced
   310   by an internal simplification procedure, which is also part of the
   311   standard Simplifier setup.
   312 
   313   \item Inject equations of a form analogous to @{prop "(x, y) = (x',
   314   y') \<equiv> x = x' \<and> y = y'"} are declared to the Simplifier and Classical
   315   Reasoner as @{attribute iff} rules.  These rules are available as
   316   @{text "t.iffs"}.
   317 
   318   \item The introduction rule for record equality analogous to @{text
   319   "x r = x r' \<Longrightarrow> y r = y r' \<dots> \<Longrightarrow> r = r'"} is declared to the Simplifier,
   320   and as the basic rule context as ``@{attribute intro}@{text "?"}''.
   321   The rule is called @{text "t.equality"}.
   322 
   323   \item Representations of arbitrary record expressions as canonical
   324   constructor terms are provided both in @{method cases} and @{method
   325   induct} format (cf.\ the generic proof methods of the same name,
   326   \secref{sec:cases-induct}).  Several variations are available, for
   327   fixed records, record schemes, more parts etc.
   328 
   329   The generic proof methods are sufficiently smart to pick the most
   330   sensible rule according to the type of the indicated record
   331   expression: users just need to apply something like ``@{text "(cases
   332   r)"}'' to a certain proof problem.
   333 
   334   \item The derived record operations @{text "t.make"}, @{text
   335   "t.fields"}, @{text "t.extend"}, @{text "t.truncate"} are \emph{not}
   336   treated automatically, but usually need to be expanded by hand,
   337   using the collective fact @{text "t.defs"}.
   338 
   339   \end{enumerate}
   340 *}
   341 
   342 
   343 section {* Datatypes \label{sec:hol-datatype} *}
   344 
   345 text {*
   346   \begin{matharray}{rcl}
   347     @{command_def (HOL) "datatype"} & : & @{text "theory \<rightarrow> theory"} \\
   348     @{command_def (HOL) "rep_datatype"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   349   \end{matharray}
   350 
   351   \begin{rail}
   352     'datatype' (dtspec + 'and')
   353     ;
   354     'rep_datatype' ('(' (name +) ')')? (term +)
   355     ;
   356 
   357     dtspec: parname? typespec mixfix? '=' (cons + '|')
   358     ;
   359     cons: name ( type * ) mixfix?
   360   \end{rail}
   361 
   362   \begin{description}
   363 
   364   \item @{command (HOL) "datatype"} defines inductive datatypes in
   365   HOL.
   366 
   367   \item @{command (HOL) "rep_datatype"} represents existing types as
   368   inductive ones, generating the standard infrastructure of derived
   369   concepts (primitive recursion etc.).
   370 
   371   \end{description}
   372 
   373   The induction and exhaustion theorems generated provide case names
   374   according to the constructors involved, while parameters are named
   375   after the types (see also \secref{sec:cases-induct}).
   376 
   377   See \cite{isabelle-HOL} for more details on datatypes, but beware of
   378   the old-style theory syntax being used there!  Apart from proper
   379   proof methods for case-analysis and induction, there are also
   380   emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL)
   381   induct_tac} available, see \secref{sec:hol-induct-tac}; these admit
   382   to refer directly to the internal structure of subgoals (including
   383   internally bound parameters).
   384 *}
   385 
   386 
   387 section {* Functorial structure of types *}
   388 
   389 text {*
   390   \begin{matharray}{rcl}
   391     @{command_def (HOL) "enriched_type"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
   392   \end{matharray}
   393 
   394   \begin{rail}
   395     'enriched_type' (prefix ':')? term
   396     ;
   397   \end{rail}
   398 
   399   \begin{description}
   400 
   401   \item @{command (HOL) "enriched_type"} allows to prove and register
   402   properties about the functorial structure of type constructors;
   403   these properties then can be used by other packages to
   404   deal with those type constructors in certain type constructions.
   405   Characteristic theorems are noted in the current local theory; by
   406   default, they are prefixed with the base name of the type constructor,
   407   an explicit prefix can be given alternatively.
   408 
   409   The given term @{text "m"} is considered as \emph{mapper} for the
   410   corresponding type constructor and must conform to the following
   411   type pattern:
   412 
   413   \begin{matharray}{lll}
   414     @{text "m"} & @{text "::"} &
   415       @{text "\<sigma>\<^isub>1 \<Rightarrow> \<dots> \<sigma>\<^isub>k \<Rightarrow> (\<^vec>\<alpha>\<^isub>n) t \<Rightarrow> (\<^vec>\<beta>\<^isub>n) t"} \\
   416   \end{matharray}
   417 
   418   \noindent where @{text t} is the type constructor, @{text
   419   "\<^vec>\<alpha>\<^isub>n"} and @{text "\<^vec>\<beta>\<^isub>n"} are distinct
   420   type variables free in the local theory and @{text "\<sigma>\<^isub>1"},
   421   \ldots, @{text "\<sigma>\<^isub>k"} is a subsequence of @{text "\<alpha>\<^isub>1 \<Rightarrow>
   422   \<beta>\<^isub>1"}, @{text "\<beta>\<^isub>1 \<Rightarrow> \<alpha>\<^isub>1"}, \ldots,
   423   @{text "\<alpha>\<^isub>n \<Rightarrow> \<beta>\<^isub>n"}, @{text "\<beta>\<^isub>n \<Rightarrow>
   424   \<alpha>\<^isub>n"}.
   425 
   426   \end{description}
   427 *}
   428 
   429 
   430 section {* Recursive functions \label{sec:recursion} *}
   431 
   432 text {*
   433   \begin{matharray}{rcl}
   434     @{command_def (HOL) "primrec"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   435     @{command_def (HOL) "fun"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   436     @{command_def (HOL) "function"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
   437     @{command_def (HOL) "termination"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
   438   \end{matharray}
   439 
   440   \begin{rail}
   441     'primrec' target? fixes 'where' equations
   442     ;
   443     ('fun' | 'function') target? functionopts? fixes \\ 'where' equations
   444     ;
   445     equations: (thmdecl? prop + '|')
   446     ;
   447     functionopts: '(' (('sequential' | 'domintros') + ',') ')'
   448     ;
   449     'termination' ( term )?
   450   \end{rail}
   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   \begin{rail}
   533     'relation' term
   534     ;
   535     'lexicographic_order' ( clasimpmod * )
   536     ;
   537     'size_change' ( orders ( clasimpmod * ) )
   538     ;
   539     orders: ( 'max' | 'min' | 'ms' ) *
   540   \end{rail}
   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   \begin{rail}
   591     'partial_function' target? '(' mode ')' fixes \\ 'where' thmdecl? prop
   592   \end{rail}
   593 
   594   \begin{description}
   595 
   596   \item @{command (HOL) "partial_function"} defines recursive
   597   functions based on fixpoints in complete partial orders. No
   598   termination proof is required from the user or constructed
   599   internally. Instead, the possibility of non-termination is modelled
   600   explicitly in the result type, which contains an explicit bottom
   601   element.
   602 
   603   Pattern matching and mutual recursion are currently not supported.
   604   Thus, the specification consists of a single function described by a
   605   single recursive equation.
   606 
   607   There are no fixed syntactic restrictions on the body of the
   608   function, but the induced functional must be provably monotonic
   609   wrt.\ the underlying order.  The monotonicitity proof is performed
   610   internally, and the definition is rejected when it fails. The proof
   611   can be influenced by declaring hints using the
   612   @{attribute (HOL) partial_function_mono} attribute.
   613 
   614   The mandatory @{text mode} argument specifies the mode of operation
   615   of the command, which directly corresponds to a complete partial
   616   order on the result type. By default, the following modes are
   617   defined:
   618 
   619   \begin{description}
   620   \item @{text option} defines functions that map into the @{type
   621   option} type. Here, the value @{term None} is used to model a
   622   non-terminating computation. Monotonicity requires that if @{term
   623   None} is returned by a recursive call, then the overall result
   624   must also be @{term None}. This is best achieved through the use of
   625   the monadic operator @{const "Option.bind"}.
   626 
   627   \item @{text tailrec} defines functions with an arbitrary result
   628   type and uses the slightly degenerated partial order where @{term
   629   "undefined"} is the bottom element.  Now, monotonicity requires that
   630   if @{term undefined} is returned by a recursive call, then the
   631   overall result must also be @{term undefined}. In practice, this is
   632   only satisfied when each recursive call is a tail call, whose result
   633   is directly returned. Thus, this mode of operation allows the
   634   definition of arbitrary tail-recursive functions.
   635   \end{description}
   636 
   637   Experienced users may define new modes by instantiating the locale
   638   @{const "partial_function_definitions"} appropriately.
   639 
   640   \item @{attribute (HOL) partial_function_mono} declares rules for
   641   use in the internal monononicity proofs of partial function
   642   definitions.
   643 
   644   \end{description}
   645 
   646 *}
   647 
   648 subsection {* Old-style recursive function definitions (TFL) *}
   649 
   650 text {*
   651   The old TFL commands @{command (HOL) "recdef"} and @{command (HOL)
   652   "recdef_tc"} for defining recursive are mostly obsolete; @{command
   653   (HOL) "function"} or @{command (HOL) "fun"} should be used instead.
   654 
   655   \begin{matharray}{rcl}
   656     @{command_def (HOL) "recdef"} & : & @{text "theory \<rightarrow> theory)"} \\
   657     @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   658   \end{matharray}
   659 
   660   \begin{rail}
   661     'recdef' ('(' 'permissive' ')')? \\ name term (prop +) hints?
   662     ;
   663     recdeftc thmdecl? tc
   664     ;
   665     hints: '(' 'hints' ( recdefmod * ) ')'
   666     ;
   667     recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf') (() | 'add' | 'del') ':' thmrefs) | clasimpmod
   668     ;
   669     tc: nameref ('(' nat ')')?
   670     ;
   671   \end{rail}
   672 
   673   \begin{description}
   674 
   675   \item @{command (HOL) "recdef"} defines general well-founded
   676   recursive functions (using the TFL package), see also
   677   \cite{isabelle-HOL}.  The ``@{text "(permissive)"}'' option tells
   678   TFL to recover from failed proof attempts, returning unfinished
   679   results.  The @{text recdef_simp}, @{text recdef_cong}, and @{text
   680   recdef_wf} hints refer to auxiliary rules to be used in the internal
   681   automated proof process of TFL.  Additional @{syntax clasimpmod}
   682   declarations (cf.\ \secref{sec:clasimp}) may be given to tune the
   683   context of the Simplifier (cf.\ \secref{sec:simplifier}) and
   684   Classical reasoner (cf.\ \secref{sec:classical}).
   685 
   686   \item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the
   687   proof for leftover termination condition number @{text i} (default
   688   1) as generated by a @{command (HOL) "recdef"} definition of
   689   constant @{text c}.
   690 
   691   Note that in most cases, @{command (HOL) "recdef"} is able to finish
   692   its internal proofs without manual intervention.
   693 
   694   \end{description}
   695 
   696   \medskip Hints for @{command (HOL) "recdef"} may be also declared
   697   globally, using the following attributes.
   698 
   699   \begin{matharray}{rcl}
   700     @{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\
   701     @{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\
   702     @{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\
   703   \end{matharray}
   704 
   705   \begin{rail}
   706     ('recdef_simp' | 'recdef_cong' | 'recdef_wf') (() | 'add' | 'del')
   707     ;
   708   \end{rail}
   709 *}
   710 
   711 
   712 section {* Inductive and coinductive definitions \label{sec:hol-inductive} *}
   713 
   714 text {*
   715   An \textbf{inductive definition} specifies the least predicate (or
   716   set) @{text R} closed under given rules: applying a rule to elements
   717   of @{text R} yields a result within @{text R}.  For example, a
   718   structural operational semantics is an inductive definition of an
   719   evaluation relation.
   720 
   721   Dually, a \textbf{coinductive definition} specifies the greatest
   722   predicate~/ set @{text R} that is consistent with given rules: every
   723   element of @{text R} can be seen as arising by applying a rule to
   724   elements of @{text R}.  An important example is using bisimulation
   725   relations to formalise equivalence of processes and infinite data
   726   structures.
   727 
   728   \medskip The HOL package is related to the ZF one, which is
   729   described in a separate paper,\footnote{It appeared in CADE
   730   \cite{paulson-CADE}; a longer version is distributed with Isabelle.}
   731   which you should refer to in case of difficulties.  The package is
   732   simpler than that of ZF thanks to implicit type-checking in HOL.
   733   The types of the (co)inductive predicates (or sets) determine the
   734   domain of the fixedpoint definition, and the package does not have
   735   to use inference rules for type-checking.
   736 
   737   \begin{matharray}{rcl}
   738     @{command_def (HOL) "inductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   739     @{command_def (HOL) "inductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   740     @{command_def (HOL) "coinductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   741     @{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   742     @{attribute_def (HOL) mono} & : & @{text attribute} \\
   743   \end{matharray}
   744 
   745   \begin{rail}
   746     ('inductive' | 'inductive_set' | 'coinductive' | 'coinductive_set') target? fixes ('for' fixes)? \\
   747     ('where' clauses)? ('monos' thmrefs)?
   748     ;
   749     clauses: (thmdecl? prop + '|')
   750     ;
   751     'mono' (() | 'add' | 'del')
   752     ;
   753   \end{rail}
   754 
   755   \begin{description}
   756 
   757   \item @{command (HOL) "inductive"} and @{command (HOL)
   758   "coinductive"} define (co)inductive predicates from the
   759   introduction rules given in the @{keyword "where"} part.  The
   760   optional @{keyword "for"} part contains a list of parameters of the
   761   (co)inductive predicates that remain fixed throughout the
   762   definition.  The optional @{keyword "monos"} section contains
   763   \emph{monotonicity theorems}, which are required for each operator
   764   applied to a recursive set in the introduction rules.  There
   765   \emph{must} be a theorem of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"},
   766   for each premise @{text "M R\<^sub>i t"} in an introduction rule!
   767 
   768   \item @{command (HOL) "inductive_set"} and @{command (HOL)
   769   "coinductive_set"} are wrappers for to the previous commands,
   770   allowing the definition of (co)inductive sets.
   771 
   772   \item @{attribute (HOL) mono} declares monotonicity rules.  These
   773   rule are involved in the automated monotonicity proof of @{command
   774   (HOL) "inductive"}.
   775 
   776   \end{description}
   777 *}
   778 
   779 
   780 subsection {* Derived rules *}
   781 
   782 text {*
   783   Each (co)inductive definition @{text R} adds definitions to the
   784   theory and also proves some theorems:
   785 
   786   \begin{description}
   787 
   788   \item @{text R.intros} is the list of introduction rules as proven
   789   theorems, for the recursive predicates (or sets).  The rules are
   790   also available individually, using the names given them in the
   791   theory file;
   792 
   793   \item @{text R.cases} is the case analysis (or elimination) rule;
   794 
   795   \item @{text R.induct} or @{text R.coinduct} is the (co)induction
   796   rule.
   797 
   798   \end{description}
   799 
   800   When several predicates @{text "R\<^sub>1, \<dots>, R\<^sub>n"} are
   801   defined simultaneously, the list of introduction rules is called
   802   @{text "R\<^sub>1_\<dots>_R\<^sub>n.intros"}, the case analysis rules are
   803   called @{text "R\<^sub>1.cases, \<dots>, R\<^sub>n.cases"}, and the list
   804   of mutual induction rules is called @{text
   805   "R\<^sub>1_\<dots>_R\<^sub>n.inducts"}.
   806 *}
   807 
   808 
   809 subsection {* Monotonicity theorems *}
   810 
   811 text {*
   812   Each theory contains a default set of theorems that are used in
   813   monotonicity proofs.  New rules can be added to this set via the
   814   @{attribute (HOL) mono} attribute.  The HOL theory @{text Inductive}
   815   shows how this is done.  In general, the following monotonicity
   816   theorems may be added:
   817 
   818   \begin{itemize}
   819 
   820   \item Theorems of the form @{text "A \<le> B \<Longrightarrow> M A \<le> M B"}, for proving
   821   monotonicity of inductive definitions whose introduction rules have
   822   premises involving terms such as @{text "M R\<^sub>i t"}.
   823 
   824   \item Monotonicity theorems for logical operators, which are of the
   825   general form @{text "(\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> (\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> \<longrightarrow> \<dots>"}.  For example, in
   826   the case of the operator @{text "\<or>"}, the corresponding theorem is
   827   \[
   828   \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"}}
   829   \]
   830 
   831   \item De Morgan style equations for reasoning about the ``polarity''
   832   of expressions, e.g.
   833   \[
   834   @{prop "\<not> \<not> P \<longleftrightarrow> P"} \qquad\qquad
   835   @{prop "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"}
   836   \]
   837 
   838   \item Equations for reducing complex operators to more primitive
   839   ones whose monotonicity can easily be proved, e.g.
   840   \[
   841   @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> \<not> P \<or> Q"} \qquad\qquad
   842   @{prop "Ball A P \<equiv> \<forall>x. x \<in> A \<longrightarrow> P x"}
   843   \]
   844 
   845   \end{itemize}
   846 
   847   %FIXME: Example of an inductive definition
   848 *}
   849 
   850 
   851 section {* Arithmetic proof support *}
   852 
   853 text {*
   854   \begin{matharray}{rcl}
   855     @{method_def (HOL) arith} & : & @{text method} \\
   856     @{attribute_def (HOL) arith} & : & @{text attribute} \\
   857     @{attribute_def (HOL) arith_split} & : & @{text attribute} \\
   858   \end{matharray}
   859 
   860   The @{method (HOL) arith} method decides linear arithmetic problems
   861   (on types @{text nat}, @{text int}, @{text real}).  Any current
   862   facts are inserted into the goal before running the procedure.
   863 
   864   The @{attribute (HOL) arith} attribute declares facts that are
   865   always supplied to the arithmetic provers implicitly.
   866 
   867   The @{attribute (HOL) arith_split} attribute declares case split
   868   rules to be expanded before @{method (HOL) arith} is invoked.
   869 
   870   Note that a simpler (but faster) arithmetic prover is
   871   already invoked by the Simplifier.
   872 *}
   873 
   874 
   875 section {* Intuitionistic proof search *}
   876 
   877 text {*
   878   \begin{matharray}{rcl}
   879     @{method_def (HOL) iprover} & : & @{text method} \\
   880   \end{matharray}
   881 
   882   \begin{rail}
   883     'iprover' ( rulemod * )
   884     ;
   885   \end{rail}
   886 
   887   The @{method (HOL) iprover} method performs intuitionistic proof
   888   search, depending on specifically declared rules from the context,
   889   or given as explicit arguments.  Chained facts are inserted into the
   890   goal before commencing proof search.
   891 
   892   Rules need to be classified as @{attribute (Pure) intro},
   893   @{attribute (Pure) elim}, or @{attribute (Pure) dest}; here the
   894   ``@{text "!"}'' indicator refers to ``safe'' rules, which may be
   895   applied aggressively (without considering back-tracking later).
   896   Rules declared with ``@{text "?"}'' are ignored in proof search (the
   897   single-step @{method rule} method still observes these).  An
   898   explicit weight annotation may be given as well; otherwise the
   899   number of rule premises will be taken into account here.
   900 *}
   901 
   902 
   903 section {* Coherent Logic *}
   904 
   905 text {*
   906   \begin{matharray}{rcl}
   907     @{method_def (HOL) "coherent"} & : & @{text method} \\
   908   \end{matharray}
   909 
   910   \begin{rail}
   911     'coherent' thmrefs?
   912     ;
   913   \end{rail}
   914 
   915   The @{method (HOL) coherent} method solves problems of
   916   \emph{Coherent Logic} \cite{Bezem-Coquand:2005}, which covers
   917   applications in confluence theory, lattice theory and projective
   918   geometry.  See @{file "~~/src/HOL/ex/Coherent.thy"} for some
   919   examples.
   920 *}
   921 
   922 
   923 section {* Checking and refuting propositions *}
   924 
   925 text {*
   926   Identifying incorrect propositions usually involves evaluation of
   927   particular assignments and systematic counter example search.  This
   928   is supported by the following commands.
   929 
   930   \begin{matharray}{rcl}
   931     @{command_def (HOL) "value"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
   932     @{command_def (HOL) "quickcheck"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
   933     @{command_def (HOL) "quickcheck_params"} & : & @{text "theory \<rightarrow> theory"}
   934   \end{matharray}
   935 
   936   \begin{rail}
   937     'value' ( ( '[' name ']' ) ? ) modes? term
   938     ;
   939 
   940     'quickcheck' ( ( '[' args ']' ) ? ) nat?
   941     ;
   942 
   943     'quickcheck_params' ( ( '[' args ']' ) ? )
   944     ;
   945 
   946     modes: '(' (name + ) ')'
   947     ;
   948 
   949     args: ( name '=' value + ',' )
   950     ;
   951   \end{rail}
   952 
   953   \begin{description}
   954 
   955   \item @{command (HOL) "value"}~@{text t} evaluates and prints a
   956     term; optionally @{text modes} can be specified, which are
   957     appended to the current print mode (see also \cite{isabelle-ref}).
   958     Internally, the evaluation is performed by registered evaluators,
   959     which are invoked sequentially until a result is returned.
   960     Alternatively a specific evaluator can be selected using square
   961     brackets; typical evaluators use the current set of code equations
   962     to normalize and include @{text simp} for fully symbolic evaluation
   963     using the simplifier, @{text nbe} for \emph{normalization by evaluation}
   964     and \emph{code} for code generation in SML.
   965 
   966   \item @{command (HOL) "quickcheck"} tests the current goal for
   967     counter examples using a series of assignments for its
   968     free variables; by default the first subgoal is tested, an other
   969     can be selected explicitly using an optional goal index.
   970     Assignments can be chosen exhausting the search space upto a given
   971     size or using a fixed number of random assignments in the search space.
   972     By default, quickcheck uses exhaustive testing.
   973     A number of configuration options are supported for
   974     @{command (HOL) "quickcheck"}, notably:
   975 
   976     \begin{description}
   977 
   978     \item[@{text tester}] specifies how to explore the search space
   979       (e.g. exhaustive or random).
   980       An unknown configuration option is treated as an argument to tester,
   981       making @{text "tester ="} optional.
   982     \item[@{text size}] specifies the maximum size of the search space
   983     for assignment values.
   984 
   985     \item[@{text eval}] takes a term or a list of terms and evaluates
   986       these terms under the variable assignment found by quickcheck.
   987 
   988     \item[@{text iterations}] sets how many sets of assignments are
   989     generated for each particular size.
   990 
   991     \item[@{text no_assms}] specifies whether assumptions in
   992     structured proofs should be ignored.
   993 
   994     \item[@{text timeout}] sets the time limit in seconds.
   995 
   996     \item[@{text default_type}] sets the type(s) generally used to
   997     instantiate type variables.
   998 
   999     \item[@{text report}] if set quickcheck reports how many tests
  1000     fulfilled the preconditions.
  1001 
  1002     \item[@{text quiet}] if not set quickcheck informs about the
  1003     current size for assignment values.
  1004 
  1005     \item[@{text expect}] can be used to check if the user's
  1006     expectation was met (@{text no_expectation}, @{text
  1007     no_counterexample}, or @{text counterexample}).
  1008 
  1009     \end{description}
  1010 
  1011     These option can be given within square brackets.
  1012 
  1013   \item @{command (HOL) "quickcheck_params"} changes quickcheck
  1014     configuration options persitently.
  1015 
  1016   \end{description}
  1017 *}
  1018 
  1019 
  1020 section {* Unstructured case analysis and induction \label{sec:hol-induct-tac} *}
  1021 
  1022 text {*
  1023   The following tools of Isabelle/HOL support cases analysis and
  1024   induction in unstructured tactic scripts; see also
  1025   \secref{sec:cases-induct} for proper Isar versions of similar ideas.
  1026 
  1027   \begin{matharray}{rcl}
  1028     @{method_def (HOL) case_tac}@{text "\<^sup>*"} & : & @{text method} \\
  1029     @{method_def (HOL) induct_tac}@{text "\<^sup>*"} & : & @{text method} \\
  1030     @{method_def (HOL) ind_cases}@{text "\<^sup>*"} & : & @{text method} \\
  1031     @{command_def (HOL) "inductive_cases"}@{text "\<^sup>*"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
  1032   \end{matharray}
  1033 
  1034   \begin{rail}
  1035     'case_tac' goalspec? term rule?
  1036     ;
  1037     'induct_tac' goalspec? (insts * 'and') rule?
  1038     ;
  1039     'ind_cases' (prop +) ('for' (name +)) ?
  1040     ;
  1041     'inductive_cases' (thmdecl? (prop +) + 'and')
  1042     ;
  1043 
  1044     rule: ('rule' ':' thmref)
  1045     ;
  1046   \end{rail}
  1047 
  1048   \begin{description}
  1049 
  1050   \item @{method (HOL) case_tac} and @{method (HOL) induct_tac} admit
  1051   to reason about inductive types.  Rules are selected according to
  1052   the declarations by the @{attribute cases} and @{attribute induct}
  1053   attributes, cf.\ \secref{sec:cases-induct}.  The @{command (HOL)
  1054   datatype} package already takes care of this.
  1055 
  1056   These unstructured tactics feature both goal addressing and dynamic
  1057   instantiation.  Note that named rule cases are \emph{not} provided
  1058   as would be by the proper @{method cases} and @{method induct} proof
  1059   methods (see \secref{sec:cases-induct}).  Unlike the @{method
  1060   induct} method, @{method induct_tac} does not handle structured rule
  1061   statements, only the compact object-logic conclusion of the subgoal
  1062   being addressed.
  1063 
  1064   \item @{method (HOL) ind_cases} and @{command (HOL)
  1065   "inductive_cases"} provide an interface to the internal @{ML_text
  1066   mk_cases} operation.  Rules are simplified in an unrestricted
  1067   forward manner.
  1068 
  1069   While @{method (HOL) ind_cases} is a proof method to apply the
  1070   result immediately as elimination rules, @{command (HOL)
  1071   "inductive_cases"} provides case split theorems at the theory level
  1072   for later use.  The @{keyword "for"} argument of the @{method (HOL)
  1073   ind_cases} method allows to specify a list of variables that should
  1074   be generalized before applying the resulting rule.
  1075 
  1076   \end{description}
  1077 *}
  1078 
  1079 
  1080 section {* Executable code *}
  1081 
  1082 text {*
  1083   Isabelle/Pure provides two generic frameworks to support code
  1084   generation from executable specifications.  Isabelle/HOL
  1085   instantiates these mechanisms in a way that is amenable to end-user
  1086   applications.
  1087 
  1088   \medskip One framework generates code from functional programs
  1089   (including overloading using type classes) to SML \cite{SML}, OCaml
  1090   \cite{OCaml}, Haskell \cite{haskell-revised-report} and Scala
  1091   \cite{scala-overview-tech-report}.
  1092   Conceptually, code generation is split up in three steps:
  1093   \emph{selection} of code theorems, \emph{translation} into an
  1094   abstract executable view and \emph{serialization} to a specific
  1095   \emph{target language}.  Inductive specifications can be executed
  1096   using the predicate compiler which operates within HOL.
  1097   See \cite{isabelle-codegen} for an introduction.
  1098 
  1099   \begin{matharray}{rcl}
  1100     @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1101     @{attribute_def (HOL) code} & : & @{text attribute} \\
  1102     @{command_def (HOL) "code_abort"} & : & @{text "theory \<rightarrow> theory"} \\
  1103     @{command_def (HOL) "code_datatype"} & : & @{text "theory \<rightarrow> theory"} \\
  1104     @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1105     @{attribute_def (HOL) code_inline} & : & @{text attribute} \\
  1106     @{attribute_def (HOL) code_post} & : & @{text attribute} \\
  1107     @{command_def (HOL) "print_codeproc"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1108     @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1109     @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1110     @{command_def (HOL) "code_const"} & : & @{text "theory \<rightarrow> theory"} \\
  1111     @{command_def (HOL) "code_type"} & : & @{text "theory \<rightarrow> theory"} \\
  1112     @{command_def (HOL) "code_class"} & : & @{text "theory \<rightarrow> theory"} \\
  1113     @{command_def (HOL) "code_instance"} & : & @{text "theory \<rightarrow> theory"} \\
  1114     @{command_def (HOL) "code_reserved"} & : & @{text "theory \<rightarrow> theory"} \\
  1115     @{command_def (HOL) "code_monad"} & : & @{text "theory \<rightarrow> theory"} \\
  1116     @{command_def (HOL) "code_include"} & : & @{text "theory \<rightarrow> theory"} \\
  1117     @{command_def (HOL) "code_modulename"} & : & @{text "theory \<rightarrow> theory"} \\
  1118     @{command_def (HOL) "code_reflect"} & : & @{text "theory \<rightarrow> theory"}
  1119   \end{matharray}
  1120 
  1121   \begin{rail}
  1122      'export_code' ( constexpr + ) \\
  1123        ( ( 'in' target ( 'module_name' string ) ? \\
  1124         ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
  1125     ;
  1126 
  1127     const: term
  1128     ;
  1129 
  1130     constexpr: ( const | 'name._' | '_' )
  1131     ;
  1132 
  1133     typeconstructor: nameref
  1134     ;
  1135 
  1136     class: nameref
  1137     ;
  1138 
  1139     target: 'SML' | 'OCaml' | 'Haskell' | 'Scala'
  1140     ;
  1141 
  1142     'code' ( 'del' | 'abstype' | 'abstract' ) ?
  1143     ;
  1144 
  1145     'code_abort' ( const + )
  1146     ;
  1147 
  1148     'code_datatype' ( const + )
  1149     ;
  1150 
  1151     'code_inline' ( 'del' ) ?
  1152     ;
  1153 
  1154     'code_post' ( 'del' ) ?
  1155     ;
  1156 
  1157     'code_thms' ( constexpr + ) ?
  1158     ;
  1159 
  1160     'code_deps' ( constexpr + ) ?
  1161     ;
  1162 
  1163     'code_const' (const + 'and') \\
  1164       ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
  1165     ;
  1166 
  1167     'code_type' (typeconstructor + 'and') \\
  1168       ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
  1169     ;
  1170 
  1171     'code_class' (class + 'and') \\
  1172       ( ( '(' target \\ ( string ? + 'and' ) ')' ) + )
  1173     ;
  1174 
  1175     'code_instance' (( typeconstructor '::' class ) + 'and') \\
  1176       ( ( '(' target ( '-' ? + 'and' ) ')' ) + )
  1177     ;
  1178 
  1179     'code_reserved' target ( string + )
  1180     ;
  1181 
  1182     'code_monad' const const target
  1183     ;
  1184 
  1185     'code_include' target ( string ( string | '-') )
  1186     ;
  1187 
  1188     'code_modulename' target ( ( string string ) + )
  1189     ;
  1190 
  1191     'code_reflect' string \\
  1192       ( 'datatypes' ( string '=' ( '_' | ( string + '|' ) + 'and' ) ) ) ? \\
  1193       ( 'functions' ( string + ) ) ? ( 'file' string ) ?
  1194     ;
  1195 
  1196     syntax: string | ( 'infix' | 'infixl' | 'infixr' ) nat string
  1197     ;
  1198 
  1199   \end{rail}
  1200 
  1201   \begin{description}
  1202 
  1203   \item @{command (HOL) "export_code"} generates code for a given list
  1204   of constants in the specified target language(s).  If no
  1205   serialization instruction is given, only abstract code is generated
  1206   internally.
  1207 
  1208   Constants may be specified by giving them literally, referring to
  1209   all executable contants within a certain theory by giving @{text
  1210   "name.*"}, or referring to \emph{all} executable constants currently
  1211   available by giving @{text "*"}.
  1212 
  1213   By default, for each involved theory one corresponding name space
  1214   module is generated.  Alternativly, a module name may be specified
  1215   after the @{keyword "module_name"} keyword; then \emph{all} code is
  1216   placed in this module.
  1217 
  1218   For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification
  1219   refers to a single file; for \emph{Haskell}, it refers to a whole
  1220   directory, where code is generated in multiple files reflecting the
  1221   module hierarchy.  Omitting the file specification denotes standard
  1222   output.
  1223 
  1224   Serializers take an optional list of arguments in parentheses.  For
  1225   \emph{SML} and \emph{OCaml}, ``@{text no_signatures}`` omits
  1226   explicit module signatures.
  1227 
  1228   For \emph{Haskell} a module name prefix may be given using the
  1229   ``@{text "root:"}'' argument; ``@{text string_classes}'' adds a
  1230   ``@{verbatim "deriving (Read, Show)"}'' clause to each appropriate
  1231   datatype declaration.
  1232 
  1233   \item @{attribute (HOL) code} explicitly selects (or with option
  1234   ``@{text "del"}'' deselects) a code equation for code generation.
  1235   Usually packages introducing code equations provide a reasonable
  1236   default setup for selection.  Variants @{text "code abstype"} and
  1237   @{text "code abstract"} declare abstract datatype certificates or
  1238   code equations on abstract datatype representations respectively.
  1239 
  1240   \item @{command (HOL) "code_abort"} declares constants which are not
  1241   required to have a definition by means of code equations; if needed
  1242   these are implemented by program abort instead.
  1243 
  1244   \item @{command (HOL) "code_datatype"} specifies a constructor set
  1245   for a logical type.
  1246 
  1247   \item @{command (HOL) "print_codesetup"} gives an overview on
  1248   selected code equations and code generator datatypes.
  1249 
  1250   \item @{attribute (HOL) code_inline} declares (or with option
  1251   ``@{text "del"}'' removes) inlining theorems which are applied as
  1252   rewrite rules to any code equation during preprocessing.
  1253 
  1254   \item @{attribute (HOL) code_post} declares (or with option ``@{text
  1255   "del"}'' removes) theorems which are applied as rewrite rules to any
  1256   result of an evaluation.
  1257 
  1258   \item @{command (HOL) "print_codeproc"} prints the setup of the code
  1259   generator preprocessor.
  1260 
  1261   \item @{command (HOL) "code_thms"} prints a list of theorems
  1262   representing the corresponding program containing all given
  1263   constants after preprocessing.
  1264 
  1265   \item @{command (HOL) "code_deps"} visualizes dependencies of
  1266   theorems representing the corresponding program containing all given
  1267   constants after preprocessing.
  1268 
  1269   \item @{command (HOL) "code_const"} associates a list of constants
  1270   with target-specific serializations; omitting a serialization
  1271   deletes an existing serialization.
  1272 
  1273   \item @{command (HOL) "code_type"} associates a list of type
  1274   constructors with target-specific serializations; omitting a
  1275   serialization deletes an existing serialization.
  1276 
  1277   \item @{command (HOL) "code_class"} associates a list of classes
  1278   with target-specific class names; omitting a serialization deletes
  1279   an existing serialization.  This applies only to \emph{Haskell}.
  1280 
  1281   \item @{command (HOL) "code_instance"} declares a list of type
  1282   constructor / class instance relations as ``already present'' for a
  1283   given target.  Omitting a ``@{text "-"}'' deletes an existing
  1284   ``already present'' declaration.  This applies only to
  1285   \emph{Haskell}.
  1286 
  1287   \item @{command (HOL) "code_reserved"} declares a list of names as
  1288   reserved for a given target, preventing it to be shadowed by any
  1289   generated code.
  1290 
  1291   \item @{command (HOL) "code_monad"} provides an auxiliary mechanism
  1292   to generate monadic code for Haskell.
  1293 
  1294   \item @{command (HOL) "code_include"} adds arbitrary named content
  1295   (``include'') to generated code.  A ``@{text "-"}'' as last argument
  1296   will remove an already added ``include''.
  1297 
  1298   \item @{command (HOL) "code_modulename"} declares aliasings from one
  1299   module name onto another.
  1300 
  1301   \item @{command (HOL) "code_reflect"} without a ``@{text "file"}''
  1302   argument compiles code into the system runtime environment and
  1303   modifies the code generator setup that future invocations of system
  1304   runtime code generation referring to one of the ``@{text
  1305   "datatypes"}'' or ``@{text "functions"}'' entities use these precompiled
  1306   entities.  With a ``@{text "file"}'' argument, the corresponding code
  1307   is generated into that specified file without modifying the code
  1308   generator setup.
  1309 
  1310   \end{description}
  1311 
  1312   The other framework generates code from both functional and
  1313   relational programs to SML.  See \cite{isabelle-HOL} for further
  1314   information (this actually covers the new-style theory format as
  1315   well).
  1316 
  1317   \begin{matharray}{rcl}
  1318     @{command_def (HOL) "code_module"} & : & @{text "theory \<rightarrow> theory"} \\
  1319     @{command_def (HOL) "code_library"} & : & @{text "theory \<rightarrow> theory"} \\
  1320     @{command_def (HOL) "consts_code"} & : & @{text "theory \<rightarrow> theory"} \\
  1321     @{command_def (HOL) "types_code"} & : & @{text "theory \<rightarrow> theory"} \\
  1322     @{attribute_def (HOL) code} & : & @{text attribute} \\
  1323   \end{matharray}
  1324 
  1325   \begin{rail}
  1326   ( 'code_module' | 'code_library' ) modespec ? name ? \\
  1327     ( 'file' name ) ? ( 'imports' ( name + ) ) ? \\
  1328     'contains' ( ( name '=' term ) + | term + )
  1329   ;
  1330 
  1331   modespec: '(' ( name * ) ')'
  1332   ;
  1333 
  1334   'consts_code' (codespec +)
  1335   ;
  1336 
  1337   codespec: const template attachment ?
  1338   ;
  1339 
  1340   'types_code' (tycodespec +)
  1341   ;
  1342 
  1343   tycodespec: name template attachment ?
  1344   ;
  1345 
  1346   const: term
  1347   ;
  1348 
  1349   template: '(' string ')'
  1350   ;
  1351 
  1352   attachment: 'attach' modespec ? verblbrace text verbrbrace
  1353   ;
  1354 
  1355   'code' (name)?
  1356   ;
  1357   \end{rail}
  1358 
  1359 *}
  1360 
  1361 
  1362 section {* Definition by specification \label{sec:hol-specification} *}
  1363 
  1364 text {*
  1365   \begin{matharray}{rcl}
  1366     @{command_def (HOL) "specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
  1367     @{command_def (HOL) "ax_specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
  1368   \end{matharray}
  1369 
  1370   \begin{rail}
  1371   ('specification' | 'ax_specification') '(' (decl +) ')' \\ (thmdecl? prop +)
  1372   ;
  1373   decl: ((name ':')? term '(' 'overloaded' ')'?)
  1374   \end{rail}
  1375 
  1376   \begin{description}
  1377 
  1378   \item @{command (HOL) "specification"}~@{text "decls \<phi>"} sets up a
  1379   goal stating the existence of terms with the properties specified to
  1380   hold for the constants given in @{text decls}.  After finishing the
  1381   proof, the theory will be augmented with definitions for the given
  1382   constants, as well as with theorems stating the properties for these
  1383   constants.
  1384 
  1385   \item @{command (HOL) "ax_specification"}~@{text "decls \<phi>"} sets up
  1386   a goal stating the existence of terms with the properties specified
  1387   to hold for the constants given in @{text decls}.  After finishing
  1388   the proof, the theory will be augmented with axioms expressing the
  1389   properties given in the first place.
  1390 
  1391   \item @{text decl} declares a constant to be defined by the
  1392   specification given.  The definition for the constant @{text c} is
  1393   bound to the name @{text c_def} unless a theorem name is given in
  1394   the declaration.  Overloaded constants should be declared as such.
  1395 
  1396   \end{description}
  1397 
  1398   Whether to use @{command (HOL) "specification"} or @{command (HOL)
  1399   "ax_specification"} is to some extent a matter of style.  @{command
  1400   (HOL) "specification"} introduces no new axioms, and so by
  1401   construction cannot introduce inconsistencies, whereas @{command
  1402   (HOL) "ax_specification"} does introduce axioms, but only after the
  1403   user has explicitly proven it to be safe.  A practical issue must be
  1404   considered, though: After introducing two constants with the same
  1405   properties using @{command (HOL) "specification"}, one can prove
  1406   that the two constants are, in fact, equal.  If this might be a
  1407   problem, one should use @{command (HOL) "ax_specification"}.
  1408 *}
  1409 
  1410 end