doc-src/IsarRef/generic.tex
author wenzelm
Fri, 17 Mar 2000 22:50:41 +0100
changeset 8507 d22fcea34cb7
parent 8483 b437907f9b26
child 8517 062e6cd78534
permissions -rw-r--r--
untag: only name arg; fixed trans att syntax; tuned case command;


\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{this}
\indexisarmeth{fold}\indexisarmeth{unfold}
\indexisarmeth{rule}\indexisarmeth{erule}
\begin{matharray}{rcl}
  - & : & \isarmeth \\
  assumption & : & \isarmeth \\
  this & : & \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.
\item [$this$] applies the current facts directly as rules.  Note that
  ``$\DOT$'' (dot) abbreviates $\BY{this}$.
\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 $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 inserted
  into the goal and subject to rewriting as well.
\item [$succeed$] yields a single (unchanged) result; it is the identity of
  the ``\texttt{,}'' method combinator (cf.\ \S\ref{sec:syn-meth}).
\item [$fail$] yields an empty result sequence; it is the identity 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}
\indexisaratt{unfold}\indexisaratt{fold}
\begin{matharray}{rcl}
  tag & : & \isaratt \\
  untag & : & \isaratt \\[0.5ex]
  OF & : & \isaratt \\
  RS & : & \isaratt \\
  COMP & : & \isaratt \\[0.5ex]
  of & : & \isaratt \\
  where & : & \isaratt \\[0.5ex]
  unfold & : & \isaratt \\
  fold & : & \isaratt \\[0.5ex]
  standard & : & \isaratt \\
  elimify & : & \isaratt \\
  export^* & : & \isaratt \\
  transfer & : & \isaratt \\[0.5ex]
\end{matharray}

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

  inst: underscore | term
  ;
\end{rail}

\begin{descr}
\item [$tag~name~args$ and $untag~name$] add and remove $tags$ of some
  theorem.  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).  The first string is considered the tag name, the rest its
  arguments.  Note that untag removes any tags of the same name.
\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 skips 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.  Arguments
  following a ``$concl\colon$'' specification refer to positions of the
  conclusion of a rule.
  
\item [$unfold~thms$ and $fold~thms$] expand and fold back again the given
  meta-level definitions throughout a rule.
 
\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')
  ;

  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
  initializes $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}$.  Note that
  ``$\FINALLY~\SHOW{}{\Var{thesis}}~\DOT$'' and
  ``$\FINALLY~\HAVE{}{\phi}~\DOT$'' are typical idioms for concluding
  calculational proofs.
  
\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{Named local contexts (cases)}\label{sec:cases}

\indexisarcmd{case}\indexisarcmd{print-cases}
\indexisaratt{case-names}\indexisaratt{params}
\begin{matharray}{rcl}
  \isarcmd{case} & : & \isartrans{proof(state)}{proof(state)} \\
  \isarcmd{print_cases} & : & \isarkeep{proof} \\[0.5ex]
  case_names & : & \isaratt \\
  params & : & \isaratt \\
\end{matharray}

Basically, Isar proof contexts are built up explicitly using commands like
$\FIXNAME$, $\ASSUMENAME$ etc.\ (see \S\ref{sec:proof-context}).  In typical
verification tasks this can become hard to manage, though.  In particular, a
large number of local contexts may emerge from case analysis or induction over
inductive sets and types.

\medskip

The $\CASENAME$ command provides a shorthand to refer to certain parts of
logical context symbolically.  Proof methods may provide an environment of
named ``cases'' of the form $c\colon \vec x, \vec \phi$.  Then the effect of
$\CASE{c}$ is exactly the same as $\FIX{\vec x}~\ASSUME{c}{\vec\phi}$.

It is important to note that $\CASENAME$ does \emph{not} provide any means to
peek at the current goal state, which is treated as strictly non-observable in
Isar!  Instead, the cases considered here usually emerge in a canonical way
from certain pieces of specification that appear in the theory somewhere else
(e.g.\ in an inductive definition, or recursive function).  See also
\S\ref{sec:induct-method} for more details of how this works in HOL.

\medskip

Named cases may be exhibited in the current proof context only if both the
proof method and the corresponding rule support this.  Case names and
parameters of basic rules may be declared by hand as well, by using
appropriate attributes.  Thus variant versions of rules that have been derived
manually may be used in advanced case analysis later.

\railalias{casenames}{case\_names}
\railterm{casenames}

\begin{rail}
  'case' nameref attributes?
  ;
  casenames (name + )
  ;
  'params' ((name * ) + 'and')
  ;
\end{rail}

\begin{descr}
\item [$\CASE{c}$] invokes a named local context $c\colon \vec x, \vec \phi$,
  as provided by an appropriate proof method (such as $cases$ and $induct$,
  see \S\ref{sec:induct-method}).  The command $\CASE{c}$ abbreviates
  $\FIX{\vec x}~\ASSUME{c}{\vec\phi}$.
\item [$\isarkeyword{print_cases}$] prints all local contexts of the current
  goal context, using Isar proof language notation.  This is a diagnostic
  command; $undo$ does not apply.
\item [$case_names~\vec c$] declares names for the local contexts of premises
  of some theorem ($\vec c$ refers to the \emph{suffix} of the list premises).
\item [$params~\vec p@1 \dots \vec p@n$] renames the innermost parameters of
  premises $1, \dots, n$ of some theorem.  An empty list of names be be given
  to skip positions, leaving the corresponding parameters unchanged.
\end{descr}


\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 \emph{definitional}
interface to type classes (cf.~\S\ref{sec:classes}).  Thus any object logic
may make use of this light-weight mechanism of 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$ to 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 $intro_classes$, 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 all class introduction rules of
  this theory.
\end{descr}

%FIXME
%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}\indexisarmeth{simp-all}
\begin{matharray}{rcl}
  simp & : & \isarmeth \\
  simp_all & : & \isarmeth \\
\end{matharray}

\railalias{simpall}{simp\_all}
\railterm{simpall}

\begin{rail}
  ('simp' | simpall) ('!' ?) (simpmod * )
  ;

  simpmod: ('add' | 'del' | 'only' | 'split' (() | 'add' | 'del') | '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, congruences, and looper tactics (including
  splits), and then behaves like \railtoken{add}.
  
  The \railtoken{split} modifiers add or delete rules for the Splitter (see
  also \cite{isabelle-ref}), the default is to add.  This works only if the
  Simplifier method has been properly setup to include the Splitter (all major
  object logics such HOL, HOLCF, FOL, ZF do this already).
  
  The \railtoken{other} modifier ignores its arguments.  Nevertheless there
  may be side-effects on the context via attributes.\footnote{This provides a
    back door for arbitrary context manipulation.}
  
\item [$simp_all$] is similar to $simp$, but acts on all goals.
\end{descr}

The $simp$ methods are based on \texttt{asm_full_simp_tac}
\cite[\S10]{isabelle-ref}, but is much better behaved in practice.  Just the
local premises of the actual goal are involved by default.  Additional facts
may be inserted via forward-chaining (using $\THEN$, $\FROMNAME$ etc.).  The
full context of assumptions is only included in the $simp!$ versions, which
should be used with some care, though.

Note that there is no separate $split$ method.  The effect of
\texttt{split_tac} can be simulated via $(simp~only\colon~split\colon~thms)$.


\subsection{Declaring rules}

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

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

\begin{descr}
\item [$simp$] adds or deletes rules from the theory or proof context (the
  default is to add).
\item [$split$] is similar to $simp$, but refers to split rules.
\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.  There are no separate options for adding or deleting
simplification rules locally.

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{intro}
\indexisarmeth{elim}\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 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 any 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 actual 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{Automated 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}.  Facts provided by forward chaining are inserted
into the goal before doing the search.  The ``!''~argument causes the full
context of assumptions to be included as well.\footnote{This is slightly less
  hazardous than for the Simplifier (see \S\ref{sec:simp}).}


\subsection{Combined automated 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' |
    ('split' (() | 'add' | 'del')) |
    (('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}.  Just note that the ones related to the
  Simplifier are prefixed by \railtoken{simp} here.
  
  Facts provided by forward chaining are inserted into the goal before doing
  the search.  The ``!''~argument causes the full context of assumptions to be
  included as well.
\end{descr}


\subsection{Declaring rules}\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 automated methods, but only in
  single-step methods such as $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: