doc-src/IsarRef/hol.tex
author wenzelm
Sun, 31 Oct 1999 15:20:35 +0100
changeset 7987 d9aef93c0e32
parent 7507 e70255cb1035
child 7990 0a604b2fc2b1
permissions -rw-r--r--
tuned;


\chapter{Isabelle/HOL Tools and Packages}\label{ch:hol-tools}

\section{Primitive types}

\indexisarcmd{typedecl}\indexisarcmd{typedef}
\begin{matharray}{rcl}
  \isarcmd{typedecl} & : & \isartrans{theory}{theory} \\
  \isarcmd{typedef} & : & \isartrans{theory}{proof(prove)} \\
\end{matharray}

\begin{rail}
  'typedecl' typespec infix? comment?
  ;
  'typedef' parname? typespec infix? \\ '=' term comment?
  ;
\end{rail}

\begin{descr}
\item [$\isarkeyword{typedecl}~(\vec\alpha)t$] is similar to the original
  $\isarkeyword{typedecl}$ of Isabelle/Pure (see \S\ref{sec:types-pure}), but
  also declares type arity $t :: (term, \dots, term) term$, making $t$ an
  actual HOL type constructor.
\item [$\isarkeyword{typedef}~(\vec\alpha)t = A$] sets up a goal stating
  non-emptiness of the set $A$.  After finishing the proof, the theory will be
  augmented by a Gordon/HOL-style type definition.  See \cite{isabelle-HOL}
  for more information.  Note that user-level theories usually do not directly
  refer to the HOL $\isarkeyword{typedef}$ primitive, but use more advanced
  packages such as $\isarkeyword{record}$ (see \S\ref{sec:record}) and
  $\isarkeyword{datatype}$ (see \S\ref{sec:datatype}).
\end{descr}


\section{Records}\label{sec:record}

%FIXME record_split method
\indexisarcmd{record}
\begin{matharray}{rcl}
  \isarcmd{record} & : & \isartrans{theory}{theory} \\
\end{matharray}

\begin{rail}
  'record' typespec '=' (type '+')? (field +)
  ;

  field: name '::' type comment?
  ;
\end{rail}

\begin{descr}
\item [$\isarkeyword{record}~(\vec\alpha)t = \tau + \vec c :: \vec\sigma$]
  defines extensible record type $(\vec\alpha)t$, derived from the optional
  parent record $\tau$ by adding new field components $\vec c :: \vec\sigma$.
  See \cite{isabelle-HOL,NaraschewskiW-TPHOLs98} for more information only
  simply-typed extensible records.
\end{descr}


\section{Datatypes}\label{sec:datatype}

\indexisarcmd{datatype}\indexisarcmd{rep-datatype}
\begin{matharray}{rcl}
  \isarcmd{datatype} & : & \isartrans{theory}{theory} \\
  \isarcmd{rep_datatype} & : & \isartrans{theory}{theory} \\
\end{matharray}

\railalias{repdatatype}{rep\_datatype}
\railterm{repdatatype}

\begin{rail}
  'datatype' (parname? typespec infix? \\ '=' (constructor + '|') + 'and')
  ;
  repdatatype (name * ) \\ 'distinct' thmrefs 'inject' thmrefs 'induction' thmrefs
  ;

  constructor: name (type * ) mixfix? comment?
  ;
\end{rail}

\begin{descr}
\item [$\isarkeyword{datatype}$] defines inductive datatypes in HOL.
\item [$\isarkeyword{rep_datatype}$] represents existing types as inductive
  ones, generating the standard infrastructure of derived concepts (primitive
  recursion etc.).
\end{descr}

See \cite{isabelle-HOL} for more details on datatypes.  Note that the theory
syntax above has been slightly simplified over the old version, usually
requiring more quotes and less parentheses.


\section{Recursive functions}

\indexisarcmd{primrec}\indexisarcmd{recdef}
\begin{matharray}{rcl}
  \isarcmd{primrec} & : & \isartrans{theory}{theory} \\
  \isarcmd{recdef} & : & \isartrans{theory}{theory} \\
%FIXME
%  \isarcmd{defer_recdef} & : & \isartrans{theory}{theory} \\
\end{matharray}

\begin{rail}
  'primrec' parname? (thmdecl? prop comment? + )
  ;
  'recdef' name term (term comment? +) \\ ('congs' thmrefs)? ('simpset' name)?
  ;
\end{rail}

\begin{descr}
\item [$\isarkeyword{primrec}$] defines primitive recursive functions over
  datatypes.
\item [$\isarkeyword{recdef}$] defines general well-founded recursive
  functions (using the TFL package).
\end{descr}

See \cite{isabelle-HOL} for more information on both mechanisms.


\section{(Co)Inductive sets}

\indexisarcmd{inductive}\indexisarcmd{coinductive}\indexisarcmd{inductive-cases}
\begin{matharray}{rcl}
  \isarcmd{inductive} & : & \isartrans{theory}{theory} \\
  \isarcmd{coinductive} & : & \isartrans{theory}{theory} \\
  \isarcmd{inductive_cases} & : & \isartrans{theory}{theory} \\
\end{matharray}

\railalias{condefs}{con\_defs}
\railalias{indcases}{inductive\_cases}
\railterm{condefs,indcases}

\begin{rail}
  ('inductive' | 'coinductive') (term comment? +) \\
    'intrs' attributes? (thmdecl? prop comment? +) \\
    'monos' thmrefs comment? \\ condefs thmrefs comment?
  ;
  indcases thmdef? nameref ':' \\ (prop +) comment?
  ;
\end{rail}

\begin{descr}
\item [$\isarkeyword{inductive}$ and $\isarkeyword{coinductive}$] define
  (co)inductive sets from the given introduction rules.
\item [$\isarkeyword{inductive_cases}$] creates simplified instances of
  elimination rules of (co)inductive sets.
\end{descr}

See \cite{isabelle-HOL} for more information.  Note that
$\isarkeyword{inductive_cases}$ corresponds to the \texttt{mk_cases} ML
function.


\section{Proof by induction}

\indexisarmeth{induct}
\begin{matharray}{rcl}
  induct & : & \isarmeth \\
\end{matharray}

The $induct$ method provides a uniform interface to induction over datatypes,
inductive sets, recursive functions etc.  Basically, it is just an interface
to the $rule$ method applied to appropriate instances of the corresponding
induction rules.

\begin{rail}
  'induct' (inst * 'and') kind?
  ;

  inst: term term?
  ;
  kind: ('type' | 'set' | 'function' | 'rule') ':' nameref
  ;
\end{rail}

\begin{descr}
\item [$induct~insts~kind$] abbreviates method $rule~R$, where $R$ is the
  induction rule specified by $kind$ and instantiated by $insts$.  The rule is
  either that of some type, set, or recursive function (defined via TFL), or
  given explicitly.
  
  The instantiation basically consists of a list of $P$ $x$ (induction
  predicate and variable) specifications, where $P$ is optional.  If $kind$ is
  omitted, the default is to pick a datatype induction rule according to the
  type of some induction variable, which may not be omitted that case.
\end{descr}


\section{Arithmetic}

\indexisarmeth{arith}
\begin{matharray}{rcl}
  arith & : & \isarmeth \\
\end{matharray}

The $arith$ method decides linear arithmetic problems (on types $nat$, $int$,
$real$).  Note that a simpler (but faster) version of arithmetic reasoning is
already performed by the Simplifier.


%%% Local Variables: 
%%% mode: latex
%%% TeX-master: "isar-ref"
%%% End: