doc-src/IsarRef/generic.tex
author wenzelm
Thu, 21 Oct 1999 18:04:07 +0200
changeset 7897 7f18f5ffbb92
parent 7526 1ea137d3b5bf
child 7905 c5f735f7428c
permissions -rw-r--r--
*** empty log message ***


\chapter{Generic Tools and Packages}\label{ch:gen-tools}

\section{Basic proof methods}\label{sec:pure-meth}

\indexisarmeth{fail}\indexisarmeth{succeed}\indexisarmeth{$-$}\indexisarmeth{assumption}
\indexisarmeth{fold}\indexisarmeth{unfold}
\indexisarmeth{rule}\indexisarmeth{erule}
\begin{matharray}{rcl}
  - & : & \isarmeth \\
  assumption & : & \isarmeth \\
  rule & : & \isarmeth \\
  erule^* & : & \isarmeth \\[0.5ex]
  fold & : & \isarmeth \\
  unfold & : & \isarmeth \\[0.5ex]
  succeed & : & \isarmeth \\
  fail & : & \isarmeth \\
\end{matharray}

\begin{rail}
  ('fold' | 'unfold' | 'rule' | 'erule') thmrefs
  ;
\end{rail}

\begin{descr}
\item [``$-$''] does nothing but insert the forward chaining facts as premises
  into the goal.  Note that command $\PROOFNAME$ without any method actually
  performs a single reduction step using the $rule$ method (see below); thus a
  plain \emph{do-nothing} proof step would be $\PROOF{-}$ rather than
  $\PROOFNAME$ alone.
\item [$assumption$] solves some goal by assumption.  Any facts given are
  guaranteed to participate in the refinement.  Note that ``$\DOT$'' (dot)
  abbreviates $\BY{assumption}$.
\item [$rule~thms$] applies some rule given as argument in backward manner;
  facts are used to reduce the rule before applying it to the goal.  Thus
  $rule$ without facts is plain \emph{introduction}, while with facts it
  becomes a (generalized) \emph{elimination}.
  
  Note that the classical reasoner introduces another version of $rule$ that
  is able to pick appropriate rules automatically, whenever explicit $thms$
  are omitted (see \S\ref{sec:classical-basic}); that method is the default
  for $\PROOFNAME$ steps.  Note that ``$\DDOT$'' (double-dot) abbreviates
  $\BY{default}$.
\item [$erule~thms$] is similar to $rule$, but applies rules by
  elim-resolution.  This is an improper method, mainly for experimentation and
  porting of old scripts.  Actual elimination proofs are usually done with
  $rule$ (single step, involving facts) or $elim$ (repeated steps, see
  \S\ref{sec:classical-basic}).
\item [$unfold~thms$ and $fold~thms$] expand and fold back again the given
  meta-level definitions throughout all goals; any facts provided are
  \emph{ignored}.
\item [$succeed$] yields a single (unchanged) result; it is the identify of
  the ``\texttt{,}'' method combinator (cf.\ \S\ref{sec:syn-meth}).
\item [$fail$] yields an empty result sequence; it is the identify of the
  ``\texttt{|}'' method combinator (cf.\ \S\ref{sec:syn-meth}).
\end{descr}


\section{Miscellaneous attributes}

\indexisaratt{tag}\indexisaratt{untag}\indexisaratt{COMP}\indexisaratt{RS}
\indexisaratt{OF}\indexisaratt{where}\indexisaratt{of}\indexisaratt{standard}
\indexisaratt{elimify}\indexisaratt{transfer}\indexisaratt{export}
\begin{matharray}{rcl}
  tag & : & \isaratt \\
  untag & : & \isaratt \\[0.5ex]
  OF & : & \isaratt \\
  RS & : & \isaratt \\
  COMP & : & \isaratt \\[0.5ex]
  of & : & \isaratt \\
  where & : & \isaratt \\[0.5ex]
  standard & : & \isaratt \\
  elimify & : & \isaratt \\
  export^* & : & \isaratt \\
  transfer & : & \isaratt \\
\end{matharray}

\begin{rail}
  ('tag' | 'untag') (nameref+)
  ;
  'OF' thmrefs
  ;
  ('RS' | 'COMP') nat? thmref
  ;
  'of' (inst * ) ('concl' ':' (inst * ))?
  ;
  'where' (name '=' term * 'and')
  ;

  inst: underscore | term
  ;
\end{rail}

\begin{descr}
\item [$tag~tags$ and $untag~tags$] add and remove $tags$ of the theorem,
  respectively.  Tags may be any list of strings that serve as comment for
  some tools (e.g.\ $\LEMMANAME$ causes the tag ``$lemma$'' to be added to the
  result).
\item [$OF~thms$, $RS~n~thm$, and $COMP~n~thm$] compose rules.  $OF$ applies
  $thms$ in parallel (cf.\ \texttt{MRS} in \cite[\S5]{isabelle-ref}, but note
  the reversed order).  Note that premises may be skipped by including $\_$
  (underscore) as argument.
  
  $RS$ resolves with the $n$-th premise of $thm$; $COMP$ is a version of $RS$
  that does not include the automatic lifting process that is normally
  intended (cf.\ \texttt{RS} and \texttt{COMP} in \cite[\S5]{isabelle-ref}).
  
\item [$of~\vec t$ and $where~\vec x = \vec t$] perform positional and named
  instantiation, respectively.  The terms given in $of$ are substituted for
  any schematic variables occurring in a theorem from left to right;
  ``\texttt{_}'' (underscore) indicates to skip a position.
 
\item [$standard$] puts a theorem into the standard form of object-rules, just
  as the ML function \texttt{standard} (see \cite[\S5]{isabelle-ref}).
  
\item [$elimify$] turns an destruction rule into an elimination, just as the
  ML function \texttt{make\_elim} (see \cite{isabelle-ref}).
  
\item [$export$] lifts a local result out of the current proof context,
  generalizing all fixed variables and discharging all assumptions.  Note that
  (partial) export is usually done automatically behind the scenes.  This
  attribute is mainly for experimentation.
  
\item [$transfer$] promotes a theorem to the current theory context, which has
  to enclose the former one.  Normally, this is done automatically when rules
  are joined by inference.

\end{descr}


\section{Calculational proof}\label{sec:calculation}

\indexisarcmd{also}\indexisarcmd{finally}\indexisaratt{trans}
\begin{matharray}{rcl}
  \isarcmd{also} & : & \isartrans{proof(state)}{proof(state)} \\
  \isarcmd{finally} & : & \isartrans{proof(state)}{proof(chain)} \\
  trans & : & \isaratt \\
\end{matharray}

Calculational proof is forward reasoning with implicit application of
transitivity rules (such those of $=$, $\le$, $<$).  Isabelle/Isar maintains
an auxiliary register $calculation$\indexisarthm{calculation} for accumulating
results obtained by transitivity composed with the current result.  Command
$\ALSO$ updates $calculation$ involving $this$, while $\FINALLY$ exhibits the
final $calculation$ by forward chaining towards the next goal statement.  Both
commands require valid current facts, i.e.\ may occur only after commands that
produce theorems such as $\ASSUMENAME$, $\NOTENAME$, or some finished proof of
$\HAVENAME$, $\SHOWNAME$ etc.

Also note that the automatic term abbreviation ``$\dots$'' has its canonical
application with calculational proofs.  It automatically refers to the
argument\footnote{The argument of a curried infix expression is its right-hand
  side.} of the preceding statement.

Isabelle/Isar calculations are implicitly subject to block structure in the
sense that new threads of calculational reasoning are commenced for any new
block (as opened by a local goal, for example).  This means that, apart from
being able to nest calculations, there is no separate \emph{begin-calculation}
command required.

\begin{rail}
  ('also' | 'finally') transrules? comment?
  ;
  'trans' (() | 'add' ':' | 'del' ':') thmrefs
  ;

  transrules: '(' thmrefs ')' interest?
  ;
\end{rail}

\begin{descr}
\item [$\ALSO~(thms)$] maintains the auxiliary $calculation$ register as
  follows.  The first occurrence of $\ALSO$ in some calculational thread
  initialises $calculation$ by $this$. Any subsequent $\ALSO$ on the same
  level of block-structure updates $calculation$ by some transitivity rule
  applied to $calculation$ and $this$ (in that order).  Transitivity rules are
  picked from the current context plus those given as $thms$ (the latter have
  precedence).
  
\item [$\FINALLY~(thms)$] maintaining $calculation$ in the same way as
  $\ALSO$, and concludes the current calculational thread.  The final result
  is exhibited as fact for forward chaining towards the next goal. Basically,
  $\FINALLY$ just abbreviates $\ALSO~\FROM{calculation}$.  Typical proof
  idioms are``$\FINALLY~\SHOW{}{\Var{thesis}}~\DOT$'' and
  ``$\FINALLY~\HAVE{}{\phi}~\DOT$''.
  
\item [$trans$] maintains the set of transitivity rules of the theory or proof
  context, by adding or deleting theorems (the default is to add).
\end{descr}

%FIXME
%See theory \texttt{HOL/Isar_examples/Group} for a simple application of
%calculations for basic equational reasoning.
%\texttt{HOL/Isar_examples/KnasterTarski} involves a few more advanced
%calculational steps in combination with natural deduction.


\section{Axiomatic Type Classes}\label{sec:axclass}

\indexisarcmd{axclass}\indexisarcmd{instance}\indexisarmeth{intro-classes}
\begin{matharray}{rcl}
  \isarcmd{axclass} & : & \isartrans{theory}{theory} \\
  \isarcmd{instance} & : & \isartrans{theory}{proof(prove)} \\
  intro_classes & : & \isarmeth \\
\end{matharray}

Axiomatic type classes are provided by Isabelle/Pure as a purely
\emph{definitional} interface to type classes (cf.~\S\ref{sec:classes}).  Thus
any object logic may make use of this light-weight mechanism for abstract
theories.  See \cite{Wenzel:1997:TPHOL} for more information.  There is also a
tutorial on \emph{Using Axiomatic Type Classes in Isabelle} that is part of
the standard Isabelle documentation.
%FIXME cite

\begin{rail}
  'axclass' classdecl (axmdecl prop comment? +)
  ;
  'instance' (nameref '<' nameref | nameref '::' simplearity) comment?
  ;
\end{rail}

\begin{descr}
\item [$\isarkeyword{axclass}~c < \vec c~axms$] defines an axiomatic type
  class as the intersection of existing classes, with additional axioms
  holding.  Class axioms may not contain more than one type variable.  The
  class axioms (with implicit sort constraints added) are bound to the given
  names.  Furthermore a class introduction rule is generated, which is
  employed by method $intro_classes$ in support instantiation proofs of this
  class.
  
\item [$\isarkeyword{instance}~c@1 < c@2$ and $\isarkeyword{instance}~t ::
  (\vec s)c$] setup up a goal stating the class relation or type arity.  The
  proof would usually proceed by the $intro_classes$ method, and then
  establish the characteristic theorems of the type classes involved.  After
  finishing the proof the theory will be augmented by a type signature
  declaration corresponding to the resulting theorem.
\item [$intro_classes$] repeatedly expands the class introduction rules of
  this theory.
\end{descr}

See theory \texttt{HOL/Isar_examples/Group} for a simple example of using
axiomatic type classes, including instantiation proofs.


\section{The Simplifier}

\subsection{Simplification methods}\label{sec:simp}

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

\begin{rail}
  'simp' (simpmod * )
  ;

  simpmod: ('add' | 'del' | 'only' | 'other') ':' thmrefs
  ;
\end{rail}

\begin{descr}
\item [$simp$] invokes Isabelle's simplifier, after modifying the context by
  adding or deleting rules as specified.  The \railtoken{only} modifier first
  removes all other rewrite rules and congruences, and then is like
  \railtoken{add}.  In contrast, \railtoken{other} ignores its arguments;
  nevertheless there may be side-effects on the context via attributes.  This
  provides a back door for arbitrary context manipulation.
  
  The $simp$ method is based on \texttt{asm_full_simp_tac} (see also
  \cite[\S10]{isabelle-ref}), but is much better behaved in practice.  Only
  the local premises of the actual goal are involved by default.  Additional
  facts may be insert via forward-chaining (using $\THEN$, $\FROMNAME$ etc.).
  The full context of assumptions is

; $simp$ removes any premises of the goal, before
  inserting the goal facts; $asm_simp$ leaves the premises.  Thus $asm_simp$
  may refer to premises that are not explicitly spelled out, potentially
  obscuring the reasoning.  The plain $simp$ method is more faithful in the
  sense that, apart from the rewrite rules of the current context, \emph{at
    most} the explicitly provided facts may participate in the simplification.
\end{descr}

\subsection{Modifying the context}

\indexisaratt{simp}
\begin{matharray}{rcl}
  simp & : & \isaratt \\
\end{matharray}

\begin{rail}
  'simp' (() | 'add' | 'del')
  ;
\end{rail}

\begin{descr}
\item [$simp$] adds or deletes rules from the theory or proof context (the
  default is to add).
\end{descr}


\subsection{Forward simplification}

\indexisaratt{simplify}\indexisaratt{asm-simplify}
\indexisaratt{full-simplify}\indexisaratt{asm-full-simplify}
\begin{matharray}{rcl}
  simplify & : & \isaratt \\
  asm_simplify & : & \isaratt \\
  full_simplify & : & \isaratt \\
  asm_full_simplify & : & \isaratt \\
\end{matharray}

These attributes provide forward rules for simplification, which should be
used only very rarely.  See the ML functions of the same name in
\cite[\S10]{isabelle-ref} for more information.


\section{The Classical Reasoner}

\subsection{Basic methods}\label{sec:classical-basic}

\indexisarmeth{rule}\indexisarmeth{default}\indexisarmeth{contradiction}
\begin{matharray}{rcl}
  rule & : & \isarmeth \\
  intro & : & \isarmeth \\
  elim & : & \isarmeth \\
  contradiction & : & \isarmeth \\
\end{matharray}

\begin{rail}
  ('rule' | 'intro' | 'elim') thmrefs
  ;
\end{rail}

\begin{descr}
\item [$rule$] as offered by the classical reasoner is a refinement over the
  primitive one (see \S\ref{sec:pure-meth}).  In the case that no rules are
  provided as arguments, it automatically determines elimination and
  introduction rules from the context (see also \S\ref{sec:classical-mod}).
  In that form it is the default method for basic proof steps, such as
  $\PROOFNAME$ and ``$\DDOT$'' (two dots).
  
\item [$intro$ and $elim$] repeatedly refine some goal by intro- or
  elim-resolution, after having inserted the facts.  Omitting the arguments
  refers to any suitable rules from the context, otherwise only the explicitly
  given ones may be applied.  The latter form admits better control of what
  actually happens, thus it is very appropriate as an initial method for
  $\PROOFNAME$ that splits up certain connectives of the goal, before entering
  the sub-proof.
  
\item [$contradiction$] solves some goal by contradiction, deriving any result
  from both $\neg A$ and $A$.  Facts, which are guaranteed to participate, may
  appear in either order.
\end{descr}


\subsection{Automatic methods}\label{sec:classical-auto}

\indexisarmeth{blast}
\indexisarmeth{fast}\indexisarmeth{best}\indexisarmeth{slow}\indexisarmeth{slow-best}
\begin{matharray}{rcl}
 blast & : & \isarmeth \\
 fast & : & \isarmeth \\
 best & : & \isarmeth \\
 slow & : & \isarmeth \\
 slow_best & : & \isarmeth \\
\end{matharray}

\railalias{slowbest}{slow\_best}
\railterm{slowbest}

\begin{rail}
  'blast' nat? (clamod * )
  ;
  ('fast' | 'best' | 'slow' | slowbest) (clamod * )
  ;

  clamod: (('intro' | 'elim' | 'dest') (() | '!' | '!!') | 'del') ':' thmrefs
  ;
\end{rail}

\begin{descr}
\item [$blast$] refers to the classical tableau prover (see \texttt{blast_tac}
  in \cite[\S11]{isabelle-ref}).  The optional argument specifies a
  user-supplied search bound (default 20).
\item [$fast$, $best$, $slow$, $slow_best$] refer to the generic classical
  reasoner (see \cite[\S11]{isabelle-ref}, tactic \texttt{fast_tac} etc).
\end{descr}

Any of above methods support additional modifiers of the context of classical
rules.  There semantics is analogous to the attributes given in
\S\ref{sec:classical-mod}.


\subsection{Combined automatic methods}

\indexisarmeth{auto}\indexisarmeth{force}
\begin{matharray}{rcl}
  force & : & \isarmeth \\
  auto & : & \isarmeth \\
\end{matharray}

\begin{rail}
  ('force' | 'auto') (clasimpmod * )
  ;

  clasimpmod: ('simp' ('add' | 'del' | 'only') | other |
    (('intro' | 'elim' | 'dest') (() | '!' | '!!') | 'del')) ':' thmrefs
\end{rail}

\begin{descr}
\item [$force$ and $auto$] provide access to Isabelle's combined
  simplification and classical reasoning tactics.  See \texttt{force_tac} and
  \texttt{auto_tac} in \cite[\S11]{isabelle-ref} for more information.  The
  modifier arguments correspond to those given in \S\ref{sec:simp} and
  \S\ref{sec:classical-auto}.  Note that the ones related to the Simplifier
  are prefixed by \railtoken{simp} here.
\end{descr}

\subsection{Modifying the context}\label{sec:classical-mod}

\indexisaratt{intro}\indexisaratt{elim}\indexisaratt{dest}
\indexisaratt{iff}\indexisaratt{delrule}
\begin{matharray}{rcl}
  intro & : & \isaratt \\
  elim & : & \isaratt \\
  dest & : & \isaratt \\
  iff & : & \isaratt \\
  delrule & : & \isaratt \\
\end{matharray}

\begin{rail}
  ('intro' | 'elim' | 'dest') (() | '!' | '!!')
  ;
\end{rail}

\begin{descr}
\item [$intro$, $elim$, $dest$] add introduction, elimination, destruct rules,
  respectively.  By default, rules are considered as \emph{safe}, while a
  single ``!'' classifies as \emph{unsafe}, and ``!!'' as \emph{extra} (i.e.\ 
  not applied in the search-oriented automatic methods, but only $rule$).
  
\item [$iff$] declares equations both as rewrite rules for the simplifier and
  classical reasoning rules.

\item [$delrule$] deletes introduction or elimination rules from the context.
  Note that destruction rules would have to be turned into elimination rules
  first, e.g.\ by using the $elimify$ attribute.
\end{descr}


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