doc-src/IsarRef/pure.tex
author kleing
Wed, 18 May 2005 06:29:42 +0200
changeset 15996 3699351b8939
parent 15744 daa84ebbdf94
child 15999 d06fc840a34c
permissions -rw-r--r--
documented new thms_containing ('rewrites' still missing)


\chapter{Basic language elements}\label{ch:pure-syntax}

Subsequently, we introduce the main part of Pure theory and proof commands,
together with fundamental proof methods and attributes.
Chapter~\ref{ch:gen-tools} describes further Isar elements provided by generic
tools and packages (such as the Simplifier) that are either part of Pure
Isabelle or pre-installed in most object logics.  Chapter~\ref{ch:logics}
refers to object-logic specific elements (mainly for HOL and ZF).

\medskip

Isar commands may be either \emph{proper} document constructors, or
\emph{improper commands}.  Some proof methods and attributes introduced later
are classified as improper as well.  Improper Isar language elements, which
are subsequently marked by ``$^*$'', are often helpful when developing proof
documents, while their use is discouraged for the final human-readable
outcome.  Typical examples are diagnostic commands that print terms or
theorems according to the current context; other commands emulate old-style
tactical theorem proving.


\section{Theory commands}

\subsection{Defining theories}\label{sec:begin-thy}

\indexisarcmd{header}\indexisarcmd{theory}\indexisarcmd{context}\indexisarcmd{end}
\begin{matharray}{rcl}
  \isarcmd{header} & : & \isarkeep{toplevel} \\
  \isarcmd{theory} & : & \isartrans{toplevel}{theory} \\
  \isarcmd{context}^* & : & \isartrans{toplevel}{theory} \\
  \isarcmd{end} & : & \isartrans{theory}{toplevel} \\
\end{matharray}

Isabelle/Isar ``new-style'' theories are either defined via theory files or
interactively.  Both theory-level specifications and proofs are handled
uniformly --- occasionally definitional mechanisms even require some explicit
proof as well.  In contrast, ``old-style'' Isabelle theories support batch
processing only, with the proof scripts collected in separate ML files.

The first ``real'' command of any theory has to be $\THEORY$, which starts a
new theory based on the merge of existing ones.  Just preceding $\THEORY$,
there may be an optional $\isarkeyword{header}$ declaration, which is relevant
to document preparation only; it acts very much like a special pre-theory
markup command (cf.\ \S\ref{sec:markup-thy} and \S\ref{sec:markup-thy}).  The
$\END$ command concludes a theory development; it has to be the very last
command of any theory file loaded in batch-mode.  The theory context may be
also changed interactively by $\CONTEXT$ without creating a new theory.

\begin{rail}
  'header' text
  ;
  'theory' name '=' (name + '+') filespecs? ':'
  ;
  'context' name
  ;

  filespecs: 'files' ((name | parname) +);
\end{rail}

\begin{descr}
\item [$\isarkeyword{header}~text$] provides plain text markup just preceding
  the formal beginning of a theory.  In actual document preparation the
  corresponding {\LaTeX} macro \verb,\isamarkupheader, may be redefined to
  produce chapter or section headings.  See also \S\ref{sec:markup-thy} and
  \S\ref{sec:markup-prf} for further markup commands.
  
\item [$\THEORY~A = B@1 + \cdots + B@n\colon$] starts a new theory $A$ based
  on the merge of existing theories $B@1, \dots, B@n$.
  
  Due to inclusion of several ancestors, the overall theory structure emerging
  in an Isabelle session forms a directed acyclic graph (DAG).  Isabelle's
  theory loader ensures that the sources contributing to the development graph
  are always up-to-date.  Changed files are automatically reloaded when
  processing theory headers interactively; batch-mode explicitly distinguishes
  \verb,update_thy, from \verb,use_thy,, see also \cite{isabelle-ref}.
  
  The optional $\isarkeyword{files}$ specification declares additional
  dependencies on ML files.  Files will be loaded immediately, unless the name
  is put in parentheses, which merely documents the dependency to be resolved
  later in the text (typically via explicit $\isarcmd{use}$ in the body text,
  see \S\ref{sec:ML}).  In reminiscence of the old-style theory system of
  Isabelle, \texttt{$A$.thy} may be also accompanied by an additional file
  \texttt{$A$.ML} consisting of ML code that is executed in the context of the
  \emph{finished} theory $A$.  That file should not be included in the
  $\isarkeyword{files}$ dependency declaration, though.
  
\item [$\CONTEXT~B$] enters an existing theory context, basically in read-only
  mode, so only a limited set of commands may be performed without destroying
  the theory.  Just as for $\THEORY$, the theory loader ensures that $B$ is
  loaded and up-to-date.
  
  This command is occasionally useful for quick interactive experiments;
  normally one should always commence a new context via $\THEORY$.
  
\item [$\END$] concludes the current theory definition or context switch.
  Note that this command cannot be undone, but the whole theory definition has
  to be retracted.

\end{descr}


\subsection{Markup commands}\label{sec:markup-thy}

\indexisarcmd{chapter}\indexisarcmd{section}\indexisarcmd{subsection}
\indexisarcmd{subsubsection}\indexisarcmd{text}\indexisarcmd{text-raw}
\begin{matharray}{rcl}
  \isarcmd{chapter} & : & \isartrans{theory}{theory} \\
  \isarcmd{section} & : & \isartrans{theory}{theory} \\
  \isarcmd{subsection} & : & \isartrans{theory}{theory} \\
  \isarcmd{subsubsection} & : & \isartrans{theory}{theory} \\
  \isarcmd{text} & : & \isartrans{theory}{theory} \\
  \isarcmd{text_raw} & : & \isartrans{theory}{theory} \\
\end{matharray}

Apart from formal comments (see \S\ref{sec:comments}), markup commands provide
a structured way to insert text into the document generated from a theory (see
\cite{isabelle-sys} for more information on Isabelle's document preparation
tools).

\railalias{textraw}{text\_raw}
\railterm{textraw}

\begin{rail}
  ('chapter' | 'section' | 'subsection' | 'subsubsection' | 'text' | textraw) text
  ;
\end{rail}

\begin{descr}
\item [$\isarkeyword{chapter}$, $\isarkeyword{section}$,
  $\isarkeyword{subsection}$, and $\isarkeyword{subsubsection}$] mark chapter
  and section headings.
\item [$\TEXT$] specifies paragraphs of plain text, including references to
  formal entities (see also \S\ref{sec:antiq} on ``antiquotations'').
\item [$\isarkeyword{text_raw}$] inserts {\LaTeX} source into the output,
  without additional markup.  Thus the full range of document manipulations
  becomes available.
\end{descr}

Any of these markup elements corresponds to a {\LaTeX} command with the name
prefixed by \verb,\isamarkup,.  For the sectioning commands this is a plain
macro with a single argument, e.g.\ \verb,\isamarkupchapter{,\dots\verb,}, for
$\isarkeyword{chapter}$.  The $\isarkeyword{text}$ markup results in a
{\LaTeX} environment \verb,\begin{isamarkuptext}, {\dots}
  \verb,\end{isamarkuptext},, while $\isarkeyword{text_raw}$ causes the text
to be inserted directly into the {\LaTeX} source.

\medskip

Additional markup commands are available for proofs (see
\S\ref{sec:markup-prf}).  Also note that the $\isarkeyword{header}$
declaration (see \S\ref{sec:begin-thy}) admits to insert section markup just
preceding the actual theory definition.


\subsection{Type classes and sorts}\label{sec:classes}

\indexisarcmd{classes}\indexisarcmd{classrel}\indexisarcmd{defaultsort}
\begin{matharray}{rcll}
  \isarcmd{classes} & : & \isartrans{theory}{theory} \\
  \isarcmd{classrel} & : & \isartrans{theory}{theory} & (axiomatic!) \\
  \isarcmd{defaultsort} & : & \isartrans{theory}{theory} \\
\end{matharray}

\begin{rail}
  'classes' (classdecl +)
  ;
  'classrel' (nameref ('<' | subseteq) nameref + 'and')
  ;
  'defaultsort' sort
  ;
\end{rail}

\begin{descr}
\item [$\isarkeyword{classes}~c \subseteq \vec c$] declares class $c$ to be a
  subclass of existing classes $\vec c$.  Cyclic class structures are ruled
  out.
\item [$\isarkeyword{classrel}~c@1 \subseteq c@2$] states subclass relations
  between existing classes $c@1$ and $c@2$.  This is done axiomatically!  The
  $\INSTANCE$ command (see \S\ref{sec:axclass}) provides a way to introduce
  proven class relations.
\item [$\isarkeyword{defaultsort}~s$] makes sort $s$ the new default sort for
  any type variables given without sort constraints.  Usually, the default
  sort would be only changed when defining a new object-logic.
\end{descr}


\subsection{Primitive types and type abbreviations}\label{sec:types-pure}

\indexisarcmd{typedecl}\indexisarcmd{types}\indexisarcmd{nonterminals}\indexisarcmd{arities}
\begin{matharray}{rcll}
  \isarcmd{types} & : & \isartrans{theory}{theory} \\
  \isarcmd{typedecl} & : & \isartrans{theory}{theory} \\
  \isarcmd{nonterminals} & : & \isartrans{theory}{theory} \\
  \isarcmd{arities} & : & \isartrans{theory}{theory} & (axiomatic!) \\
\end{matharray}

\begin{rail}
  'types' (typespec '=' type infix? +)
  ;
  'typedecl' typespec infix?
  ;
  'nonterminals' (name +)
  ;
  'arities' (nameref '::' arity +)
  ;
\end{rail}

\begin{descr}

\item [$\TYPES~(\vec\alpha)t = \tau$] introduces \emph{type synonym}
  $(\vec\alpha)t$ for existing type $\tau$.  Unlike actual type definitions,
  as are available in Isabelle/HOL for example, type synonyms are just purely
  syntactic abbreviations without any logical significance.  Internally, type
  synonyms are fully expanded.
  
\item [$\isarkeyword{typedecl}~(\vec\alpha)t$] declares a new type constructor
  $t$, intended as an actual logical type.  Note that the Isabelle/HOL
  object-logic overrides $\isarkeyword{typedecl}$ by its own version
  (\S\ref{sec:hol-typedef}).

\item [$\isarkeyword{nonterminals}~\vec c$] declares $0$-ary type constructors
  $\vec c$ to act as purely syntactic types, i.e.\ nonterminal symbols of
  Isabelle's inner syntax of terms or types.

\item [$\isarkeyword{arities}~t::(\vec s)s$] augments Isabelle's order-sorted
  signature of types by new type constructor arities.  This is done
  axiomatically!  The $\INSTANCE$ command (see \S\ref{sec:axclass}) provides a
  way to introduce proven type arities.

\end{descr}


\subsection{Constants and simple definitions}\label{sec:consts}

\indexisarcmd{consts}\indexisarcmd{defs}\indexisarcmd{constdefs}\indexoutertoken{constdecl}
\begin{matharray}{rcl}
  \isarcmd{consts} & : & \isartrans{theory}{theory} \\
  \isarcmd{defs} & : & \isartrans{theory}{theory} \\
  \isarcmd{constdefs} & : & \isartrans{theory}{theory} \\
\end{matharray}

\begin{rail}
  'consts' ((name '::' type mixfix?) +)
  ;
  'defs' ('(' 'overloaded' ')')? (axmdecl prop +)
  ;
\end{rail}

\begin{rail}
  'constdefs' structs? (constdecl? constdef +)
  ;

  structs: '(' 'structure' (vars + 'and') ')'
  ;
  constdecl: (name '::' type) mixfix | (name '::' type) | name 'where' | mixfix
  ;
  constdef: thmdecl? prop
  ;
\end{rail}

\begin{descr}
\item [$\CONSTS~c::\sigma$] declares constant $c$ to have any instance of type
  scheme $\sigma$.  The optional mixfix annotations may attach concrete syntax
  to the constants declared.

\item [$\DEFS~name: eqn$] introduces $eqn$ as a definitional axiom for some
  existing constant.  See \cite[\S6]{isabelle-ref} for more details on the
  form of equations admitted as constant definitions.
  
  The $(overloaded)$ option declares definitions to be potentially overloaded.
  Unless this option is given, a warning message would be issued for any
  definitional equation with a more special type than that of the
  corresponding constant declaration.
  
\item [$\CONSTDEFS$] provides a streamlined combination of constants
  declarations and definitions: type-inference takes care of the most general
  typing of the given specification (the optional type constraint may refer to
  type-inference dummies ``$_$'' as usual).  The resulting type declaration
  needs to agree with that of the specification; overloading is \emph{not}
  supported here!
  
  The constant name may be omitted altogether, if neither type nor syntax
  declarations are given.  The canonical name of the definitional axiom for
  constant $c$ will be $c_def$, unless specified otherwise.  Also note that
  the given list of specifications is processed in a strictly sequential
  manner, with type-checking being performed independently.
  
  An optional initial context of $(structure)$ declarations admits use of
  indexed syntax, using the special symbol \verb,\<index>, (printed as
  ``\i'').  The latter concept is particularly useful with locales (see also
  \S\ref{sec:locale}).
\end{descr}


\subsection{Syntax and translations}\label{sec:syn-trans}

\indexisarcmd{syntax}\indexisarcmd{no-syntax}\indexisarcmd{translations}
\begin{matharray}{rcl}
  \isarcmd{syntax} & : & \isartrans{theory}{theory} \\
  \isarcmd{no_syntax} & : & \isartrans{theory}{theory} \\
  \isarcmd{translations} & : & \isartrans{theory}{theory} \\
\end{matharray}

\railalias{rightleftharpoons}{\isasymrightleftharpoons}
\railterm{rightleftharpoons}

\railalias{rightharpoonup}{\isasymrightharpoonup}
\railterm{rightharpoonup}

\railalias{leftharpoondown}{\isasymleftharpoondown}
\railterm{leftharpoondown}

\railalias{nosyntax}{no\_syntax}
\railterm{nosyntax}

\begin{rail}
  ('syntax' | nosyntax) mode? (constdecl +)
  ;
  'translations' (transpat ('==' | '=>' | '<=' | rightleftharpoons | rightharpoonup | leftharpoondown) transpat +)
  ;

  mode: ('(' ( name | 'output' | name 'output' ) ')')
  ;
  transpat: ('(' nameref ')')? string
  ;
\end{rail}

\begin{descr}
  
\item [$\isarkeyword{syntax}~(mode)~decls$] is similar to $\CONSTS~decls$,
  except that the actual logical signature extension is omitted.  Thus the
  context free grammar of Isabelle's inner syntax may be augmented in
  arbitrary ways, independently of the logic.  The $mode$ argument refers to
  the print mode that the grammar rules belong; unless the
  $\isarkeyword{output}$ indicator is given, all productions are added both to
  the input and output grammar.
  
\item [$\isarkeyword{no_syntax}~(mode)~decls$] removes grammar declarations
  (and translations) resulting from $decls$, which are interpreted in the same
  manner as for $\isarkeyword{syntax}$ above.
  
\item [$\isarkeyword{translations}~rules$] specifies syntactic translation
  rules (i.e.\ macros): parse~/ print rules (\isasymrightleftharpoons), parse
  rules (\isasymrightharpoonup), or print rules (\isasymleftharpoondown).
  Translation patterns may be prefixed by the syntactic category to be used
  for parsing; the default is $logic$.
\end{descr}


\subsection{Axioms and theorems}\label{sec:axms-thms}

\indexisarcmd{axioms}\indexisarcmd{lemmas}\indexisarcmd{theorems}
\begin{matharray}{rcll}
  \isarcmd{axioms} & : & \isartrans{theory}{theory} & (axiomatic!) \\
  \isarcmd{lemmas} & : & \isartrans{theory}{theory} \\
  \isarcmd{theorems} & : & \isartrans{theory}{theory} \\
\end{matharray}

\begin{rail}
  'axioms' (axmdecl prop +)
  ;
  ('lemmas' | 'theorems') locale? (thmdef? thmrefs + 'and')
  ;
\end{rail}

\begin{descr}
  
\item [$\isarkeyword{axioms}~a: \phi$] introduces arbitrary statements as
  axioms of the meta-logic.  In fact, axioms are ``axiomatic theorems'', and
  may be referred later just as any other theorem.
  
  Axioms are usually only introduced when declaring new logical systems.
  Everyday work is typically done the hard way, with proper definitions and
  proven theorems.
  
\item [$\isarkeyword{lemmas}~a = \vec b$] retrieves and stores existing facts
  in the theory context, or the specified locale (see also
  \S\ref{sec:locale}).  Typical applications would also involve attributes, to
  declare Simplifier rules, for example.
  
\item [$\isarkeyword{theorems}$] is essentially the same as
  $\isarkeyword{lemmas}$, but marks the result as a different kind of facts.

\end{descr}


\subsection{Name spaces}

\indexisarcmd{global}\indexisarcmd{local}\indexisarcmd{hide}
\begin{matharray}{rcl}
  \isarcmd{global} & : & \isartrans{theory}{theory} \\
  \isarcmd{local} & : & \isartrans{theory}{theory} \\
  \isarcmd{hide} & : & \isartrans{theory}{theory} \\
\end{matharray}

\begin{rail}
  'hide' name (nameref + )
  ;
\end{rail}

Isabelle organizes any kind of name declarations (of types, constants,
theorems etc.) by separate hierarchically structured name spaces.  Normally
the user does not have to control the behavior of name spaces by hand, yet the
following commands provide some way to do so.

\begin{descr}
\item [$\isarkeyword{global}$ and $\isarkeyword{local}$] change the current
  name declaration mode.  Initially, theories start in $\isarkeyword{local}$
  mode, causing all names to be automatically qualified by the theory name.
  Changing this to $\isarkeyword{global}$ causes all names to be declared
  without the theory prefix, until $\isarkeyword{local}$ is declared again.
  
  Note that global names are prone to get hidden accidently later, when
  qualified names of the same base name are introduced.
  
\item [$\isarkeyword{hide}~space~names$] removes declarations from a given
  name space (which may be $class$, $type$, or $const$).  Hidden objects
  remain valid within the logic, but are inaccessible from user input.  In
  output, the special qualifier ``$\mathord?\mathord?$'' is prefixed to the
  full internal name.  Unqualified (global) names may not be hidden.
\end{descr}


\subsection{Incorporating ML code}\label{sec:ML}

\indexisarcmd{use}\indexisarcmd{ML}\indexisarcmd{ML-command}
\indexisarcmd{ML-setup}\indexisarcmd{setup}
\indexisarcmd{method-setup}
\begin{matharray}{rcl}
  \isarcmd{use} & : & \isartrans{\cdot}{\cdot} \\
  \isarcmd{ML} & : & \isartrans{\cdot}{\cdot} \\
  \isarcmd{ML_command} & : & \isartrans{\cdot}{\cdot} \\
  \isarcmd{ML_setup} & : & \isartrans{theory}{theory} \\
  \isarcmd{setup} & : & \isartrans{theory}{theory} \\
  \isarcmd{method_setup} & : & \isartrans{theory}{theory} \\
\end{matharray}

\railalias{MLsetup}{ML\_setup}
\railterm{MLsetup}

\railalias{methodsetup}{method\_setup}
\railterm{methodsetup}

\railalias{MLcommand}{ML\_command}
\railterm{MLcommand}

\begin{rail}
  'use' name
  ;
  ('ML' | MLcommand | MLsetup | 'setup') text
  ;
  methodsetup name '=' text text
  ;
\end{rail}

\begin{descr}
\item [$\isarkeyword{use}~file$] reads and executes ML commands from $file$.
  The current theory context (if present) is passed down to the ML session,
  but may not be modified.  Furthermore, the file name is checked with the
  $\isarkeyword{files}$ dependency declaration given in the theory header (see
  also \S\ref{sec:begin-thy}).
  
\item [$\isarkeyword{ML}~text$ and $\isarkeyword{ML_command}~text$] execute ML
  commands from $text$.  The theory context is passed in the same way as for
  $\isarkeyword{use}$, but may not be changed.  Note that the output of
  $\isarkeyword{ML_command}$ is less verbose than plain $\isarkeyword{ML}$.
  
\item [$\isarkeyword{ML_setup}~text$] executes ML commands from $text$.  The
  theory context is passed down to the ML session, and fetched back
  afterwards.  Thus $text$ may actually change the theory as a side effect.
  
\item [$\isarkeyword{setup}~text$] changes the current theory context by
  applying $text$, which refers to an ML expression of type
  \texttt{(theory~->~theory)~list}.  The $\isarkeyword{setup}$ command is the
  canonical way to initialize any object-logic specific tools and packages
  written in ML.
  
\item [$\isarkeyword{method_setup}~name = text~description$] defines a proof
  method in the current theory.  The given $text$ has to be an ML expression
  of type \texttt{Args.src -> Proof.context -> Proof.method}.  Parsing
  concrete method syntax from \texttt{Args.src} input can be quite tedious in
  general.  The following simple examples are for methods without any explicit
  arguments, or a list of theorems, respectively.

{\footnotesize
\begin{verbatim}
 Method.no_args (Method.METHOD (fn facts => foobar_tac))
 Method.thms_args (fn thms => Method.METHOD (fn facts => foobar_tac))
 Method.ctxt_args (fn ctxt => Method.METHOD (fn facts => foobar_tac))
 Method.thms_ctxt_args (fn thms => fn ctxt =>
    Method.METHOD (fn facts => foobar_tac))
\end{verbatim}
}

Note that mere tactic emulations may ignore the \texttt{facts} parameter
above.  Proper proof methods would do something appropriate with the list of
current facts, though.  Single-rule methods usually do strict forward-chaining
(e.g.\ by using \texttt{Method.multi_resolves}), while automatic ones just
insert the facts using \texttt{Method.insert_tac} before applying the main
tactic.
\end{descr}


\subsection{Syntax translation functions}

\indexisarcmd{parse-ast-translation}\indexisarcmd{parse-translation}
\indexisarcmd{print-translation}\indexisarcmd{typed-print-translation}
\indexisarcmd{print-ast-translation}\indexisarcmd{token-translation}
\begin{matharray}{rcl}
  \isarcmd{parse_ast_translation} & : & \isartrans{theory}{theory} \\
  \isarcmd{parse_translation} & : & \isartrans{theory}{theory} \\
  \isarcmd{print_translation} & : & \isartrans{theory}{theory} \\
  \isarcmd{typed_print_translation} & : & \isartrans{theory}{theory} \\
  \isarcmd{print_ast_translation} & : & \isartrans{theory}{theory} \\
  \isarcmd{token_translation} & : & \isartrans{theory}{theory} \\
\end{matharray}

\railalias{parseasttranslation}{parse\_ast\_translation}
\railterm{parseasttranslation}

\railalias{parsetranslation}{parse\_translation}
\railterm{parsetranslation}

\railalias{printtranslation}{print\_translation}
\railterm{printtranslation}

\railalias{typedprinttranslation}{typed\_print\_translation}
\railterm{typedprinttranslation}

\railalias{printasttranslation}{print\_ast\_translation}
\railterm{printasttranslation}

\railalias{tokentranslation}{token\_translation}
\railterm{tokentranslation}

\begin{rail}
  ( parseasttranslation | parsetranslation | printtranslation | typedprinttranslation |
  printasttranslation ) ('(' 'advanced' ')')? text;

  tokentranslation text
\end{rail}

Syntax translation functions written in ML admit almost arbitrary
manipulations of Isabelle's inner syntax.  Any of the above commands have a
single \railqtok{text} argument that refers to an ML expression of appropriate
type, which are as follows by default:

\begin{ttbox}
val parse_ast_translation   : (string * (ast list -> ast)) list
val parse_translation       : (string * (term list -> term)) list
val print_translation       : (string * (term list -> term)) list
val typed_print_translation :
  (string * (bool -> typ -> term list -> term)) list
val print_ast_translation   : (string * (ast list -> ast)) list
val token_translation       :
  (string * string * (string -> string * real)) list
\end{ttbox}

In case that the $(advanced)$ option is given, the corresponding translation
functions may depend on the signature of the current theory context.  This
allows to implement advanced syntax mechanisms, as translations functions may
refer to specific theory declarations and auxiliary data.

See also \cite[\S8]{isabelle-ref} for more information on the general concept
of syntax transformations in Isabelle.

\begin{ttbox}
val parse_ast_translation:
  (string * (Sign.sg -> ast list -> ast)) list
val parse_translation:
  (string * (Sign.sg -> term list -> term)) list
val print_translation:
  (string * (Sign.sg -> term list -> term)) list
val typed_print_translation:
  (string * (Sign.sg -> bool -> typ -> term list -> term)) list
val print_ast_translation:
  (string * (Sign.sg -> ast list -> ast)) list
\end{ttbox}


\subsection{Oracles}

\indexisarcmd{oracle}
\begin{matharray}{rcl}
  \isarcmd{oracle} & : & \isartrans{theory}{theory} \\
\end{matharray}

Oracles provide an interface to external reasoning systems, without giving up
control completely --- each theorem carries a derivation object recording any
oracle invocation.  See \cite[\S6]{isabelle-ref} for more information.

\begin{rail}
  'oracle' name '=' text
  ;
\end{rail}

\begin{descr}
\item [$\isarkeyword{oracle}~name=text$] declares oracle $name$ to be ML
  function $text$, which has to be of type
  \texttt{Sign.sg~*~Object.T~->~term}.
\end{descr}


\section{Proof commands}

Proof commands perform transitions of Isar/VM machine configurations, which
are block-structured, consisting of a stack of nodes with three main
components: logical proof context, current facts, and open goals.  Isar/VM
transitions are \emph{typed} according to the following three different modes
of operation:
\begin{descr}
\item [$proof(prove)$] means that a new goal has just been stated that is now
  to be \emph{proven}; the next command may refine it by some proof method,
  and enter a sub-proof to establish the actual result.
\item [$proof(state)$] is like a nested theory mode: the context may be
  augmented by \emph{stating} additional assumptions, intermediate results
  etc.
\item [$proof(chain)$] is intermediate between $proof(state)$ and
  $proof(prove)$: existing facts (i.e.\ the contents of the special ``$this$''
  register) have been just picked up in order to be used when refining the
  goal claimed next.
\end{descr}

The proof mode indicator may be read as a verb telling the writer what kind of
operation may be performed next.  The corresponding typings of proof commands
restricts the shape of well-formed proof texts to particular command
sequences.  So dynamic arrangements of commands eventually turn out as static
texts of a certain structure.  Appendix~\ref{ap:refcard} gives a simplified
grammar of the overall (extensible) language emerging that way.


\subsection{Markup commands}\label{sec:markup-prf}

\indexisarcmd{sect}\indexisarcmd{subsect}\indexisarcmd{subsubsect}
\indexisarcmd{txt}\indexisarcmd{txt-raw}
\begin{matharray}{rcl}
  \isarcmd{sect} & : & \isartrans{proof}{proof} \\
  \isarcmd{subsect} & : & \isartrans{proof}{proof} \\
  \isarcmd{subsubsect} & : & \isartrans{proof}{proof} \\
  \isarcmd{txt} & : & \isartrans{proof}{proof} \\
  \isarcmd{txt_raw} & : & \isartrans{proof}{proof} \\
\end{matharray}

These markup commands for proof mode closely correspond to the ones of theory
mode (see \S\ref{sec:markup-thy}).

\railalias{txtraw}{txt\_raw}
\railterm{txtraw}

\begin{rail}
  ('sect' | 'subsect' | 'subsubsect' | 'txt' | txtraw) text
  ;
\end{rail}


\subsection{Context elements}\label{sec:proof-context}

\indexisarcmd{fix}\indexisarcmd{assume}\indexisarcmd{presume}\indexisarcmd{def}
\begin{matharray}{rcl}
  \isarcmd{fix} & : & \isartrans{proof(state)}{proof(state)} \\
  \isarcmd{assume} & : & \isartrans{proof(state)}{proof(state)} \\
  \isarcmd{presume} & : & \isartrans{proof(state)}{proof(state)} \\
  \isarcmd{def} & : & \isartrans{proof(state)}{proof(state)} \\
\end{matharray}

The logical proof context consists of fixed variables and assumptions.  The
former closely correspond to Skolem constants, or meta-level universal
quantification as provided by the Isabelle/Pure logical framework.
Introducing some \emph{arbitrary, but fixed} variable via ``$\FIX x$'' results
in a local value that may be used in the subsequent proof as any other
variable or constant.  Furthermore, any result $\edrv \phi[x]$ exported from
the context will be universally closed wrt.\ $x$ at the outermost level:
$\edrv \All x \phi$ (this is expressed using Isabelle's meta-variables).

Similarly, introducing some assumption $\chi$ has two effects.  On the one
hand, a local theorem is created that may be used as a fact in subsequent
proof steps.  On the other hand, any result $\chi \drv \phi$ exported from the
context becomes conditional wrt.\ the assumption: $\edrv \chi \Imp \phi$.
Thus, solving an enclosing goal using such a result would basically introduce
a new subgoal stemming from the assumption.  How this situation is handled
depends on the actual version of assumption command used: while $\ASSUMENAME$
insists on solving the subgoal by unification with some premise of the goal,
$\PRESUMENAME$ leaves the subgoal unchanged in order to be proved later by the
user.

Local definitions, introduced by ``$\DEF{}{x \equiv t}$'', are achieved by
combining ``$\FIX x$'' with another version of assumption that causes any
hypothetical equation $x \equiv t$ to be eliminated by the reflexivity rule.
Thus, exporting some result $x \equiv t \drv \phi[x]$ yields $\edrv \phi[t]$.

\railalias{equiv}{\isasymequiv}
\railterm{equiv}

\begin{rail}
  'fix' (vars + 'and')
  ;
  ('assume' | 'presume') (props + 'and')
  ;
  'def' thmdecl? \\ name ('==' | equiv) term termpat?
  ;
\end{rail}

\begin{descr}
  
\item [$\FIX{\vec x}$] introduces local \emph{arbitrary, but fixed} variables
  $\vec x$.
  
\item [$\ASSUME{a}{\vec\phi}$ and $\PRESUME{a}{\vec\phi}$] introduce local
  theorems $\vec\phi$ by assumption.  Subsequent results applied to an
  enclosing goal (e.g.\ by $\SHOWNAME$) are handled as follows: $\ASSUMENAME$
  expects to be able to unify with existing premises in the goal, while
  $\PRESUMENAME$ leaves $\vec\phi$ as new subgoals.
  
  Several lists of assumptions may be given (separated by
  $\isarkeyword{and}$); the resulting list of current facts consists of all of
  these concatenated.
  
\item [$\DEF{a}{x \equiv t}$] introduces a local (non-polymorphic) definition.
  In results exported from the context, $x$ is replaced by $t$.  Basically,
  ``$\DEF{}{x \equiv t}$'' abbreviates ``$\FIX{x}~\ASSUME{}{x \equiv t}$'',
  with the resulting hypothetical equation solved by reflexivity.
  
  The default name for the definitional equation is $x_def$.

\end{descr}

The special name $prems$\indexisarthm{prems} refers to all assumptions of the
current context as a list of theorems.


\subsection{Facts and forward chaining}

\indexisarcmd{note}\indexisarcmd{then}\indexisarcmd{from}\indexisarcmd{with}
\indexisarcmd{using}
\begin{matharray}{rcl}
  \isarcmd{note} & : & \isartrans{proof(state)}{proof(state)} \\
  \isarcmd{then} & : & \isartrans{proof(state)}{proof(chain)} \\
  \isarcmd{from} & : & \isartrans{proof(state)}{proof(chain)} \\
  \isarcmd{with} & : & \isartrans{proof(state)}{proof(chain)} \\
  \isarcmd{using} & : & \isartrans{proof(prove)}{proof(prove)} \\
\end{matharray}

New facts are established either by assumption or proof of local statements.
Any fact will usually be involved in further proofs, either as explicit
arguments of proof methods, or when forward chaining towards the next goal via
$\THEN$ (and variants); $\FROMNAME$ and $\WITHNAME$ are composite forms
involving $\NOTENAME$.  The $\USINGNAME$ elements augments the collection of
used facts \emph{after} a goal has been stated.  Note that the special theorem
name $this$\indexisarthm{this} refers to the most recently established facts,
but only \emph{before} issuing a follow-up claim.

\begin{rail}
  'note' (thmdef? thmrefs + 'and')
  ;
  ('from' | 'with' | 'using') (thmrefs + 'and')
  ;
\end{rail}

\begin{descr}

\item [$\NOTE{a}{\vec b}$] recalls existing facts $\vec b$, binding the result
  as $a$.  Note that attributes may be involved as well, both on the left and
  right hand sides.

\item [$\THEN$] indicates forward chaining by the current facts in order to
  establish the goal to be claimed next.  The initial proof method invoked to
  refine that will be offered the facts to do ``anything appropriate'' (see
  also \S\ref{sec:proof-steps}).  For example, method $rule$ (see
  \S\ref{sec:pure-meth-att}) would typically do an elimination rather than an
  introduction.  Automatic methods usually insert the facts into the goal
  state before operation.  This provides a simple scheme to control relevance
  of facts in automated proof search.
  
\item [$\FROM{\vec b}$] abbreviates ``$\NOTE{}{\vec b}~\THEN$''; thus $\THEN$
  is equivalent to ``$\FROM{this}$''.
  
\item [$\WITH{\vec b}$] abbreviates ``$\FROM{\vec b~\AND~this}$''; thus the
  forward chaining is from earlier facts together with the current ones.
  
\item [$\USING{\vec b}$] augments the facts being currently indicated for use
  by a subsequent refinement step (such as $\APPLYNAME$ or $\PROOFNAME$).

\end{descr}

Forward chaining with an empty list of theorems is the same as not chaining at
all.  Thus ``$\FROM{nothing}$'' has no effect apart from entering
$prove(chain)$ mode, since $nothing$\indexisarthm{nothing} is bound to the
empty list of theorems.

Basic proof methods (such as $rule$) expect multiple facts to be given in
their proper order, corresponding to a prefix of the premises of the rule
involved.  Note that positions may be easily skipped using something like
$\FROM{\Text{\texttt{_}}~a~b}$, for example.  This involves the trivial rule
$\PROP\psi \Imp \PROP\psi$, which happens to be bound in Isabelle/Pure as
``\texttt{_}'' (underscore).\indexisarthm{_@\texttt{_}}

Automated methods (such as $simp$ or $auto$) just insert any given facts
before their usual operation.  Depending on the kind of procedure involved,
the order of facts is less significant here.


\subsection{Goal statements}\label{sec:goals}

\indexisarcmd{lemma}\indexisarcmd{theorem}\indexisarcmd{corollary}
\indexisarcmd{have}\indexisarcmd{show}\indexisarcmd{hence}\indexisarcmd{thus}
\begin{matharray}{rcl}
  \isarcmd{lemma} & : & \isartrans{theory}{proof(prove)} \\
  \isarcmd{theorem} & : & \isartrans{theory}{proof(prove)} \\
  \isarcmd{corollary} & : & \isartrans{theory}{proof(prove)} \\
  \isarcmd{have} & : & \isartrans{proof(state) ~|~ proof(chain)}{proof(prove)} \\
  \isarcmd{show} & : & \isartrans{proof(state) ~|~ proof(chain)}{proof(prove)} \\
  \isarcmd{hence} & : & \isartrans{proof(state)}{proof(prove)} \\
  \isarcmd{thus} & : & \isartrans{proof(state)}{proof(prove)} \\
\end{matharray}

From a theory context, proof mode is entered by an initial goal command such
as $\LEMMANAME$, $\THEOREMNAME$, or $\COROLLARYNAME$.  Within a proof, new
claims may be introduced locally as well; four variants are available here to
indicate whether forward chaining of facts should be performed initially (via
$\THEN$), and whether the final result is meant to solve some pending goal.

Goals may consist of multiple statements, resulting in a list of facts
eventually.  A pending multi-goal is internally represented as a meta-level
conjunction (printed as \verb,&&,), which is usually split into the
corresponding number of sub-goals prior to an initial method application, via
$\PROOFNAME$ (\S\ref{sec:proof-steps}) or $\APPLYNAME$
(\S\ref{sec:tactic-commands}).  The $induct$ method covered in
\S\ref{sec:cases-induct-meth} acts on multiple claims simultaneously.

Claims at the theory level may be either in short or long form.  A short goal
merely consists of several simultaneous propositions (often just one).  A long
goal includes an explicit context specification for the subsequent
conclusions, involving local parameters; here the role of each part of the
statement is explicitly marked by separate keywords (see also
\S\ref{sec:locale}).

\begin{rail}
  ('lemma' | 'theorem' | 'corollary') locale? (goal | longgoal)
  ;
  ('have' | 'show' | 'hence' | 'thus') goal
  ;
  
  goal: (props + 'and')
  ;
  longgoal: thmdecl? (contextelem *) 'shows' goal
  ;
\end{rail}

\begin{descr}
  
\item [$\LEMMA{a}{\vec\phi}$] enters proof mode with $\vec\phi$ as main goal,
  eventually resulting in some fact $\turn \vec\phi$ to be put back into the
  theory context, or into the specified locale (cf.\ \S\ref{sec:locale}).  An
  additional \railnonterm{context} specification may build up an initial proof
  context for the subsequent claim; this includes local definitions and syntax
  as well, see the definition of $contextelem$ in \S\ref{sec:locale}.
  
\item [$\THEOREM{a}{\vec\phi}$ and $\COROLLARY{a}{\vec\phi}$] are essentially
  the same as $\LEMMA{a}{\vec\phi}$, but the facts are internally marked as
  being of a different kind.  This discrimination acts like a formal comment.
  
\item [$\HAVE{a}{\vec\phi}$] claims a local goal, eventually resulting in a
  fact within the current logical context.  This operation is completely
  independent of any pending sub-goals of an enclosing goal statements, so
  $\HAVENAME$ may be freely used for experimental exploration of potential
  results within a proof body.
  
\item [$\SHOW{a}{\vec\phi}$] is like $\HAVE{a}{\vec\phi}$ plus a second stage
  to refine some pending sub-goal for each one of the finished result, after
  having been exported into the corresponding context (at the head of the
  sub-proof of this $\SHOWNAME$ command).
  
  To accommodate interactive debugging, resulting rules are printed before
  being applied internally.  Even more, interactive execution of $\SHOWNAME$
  predicts potential failure and displays the resulting error as a warning
  beforehand.  Watch out for the following message:

  \begin{ttbox}
  Problem! Local statement will fail to solve any pending goal
  \end{ttbox}
  
\item [$\HENCENAME$] abbreviates ``$\THEN~\HAVENAME$'', i.e.\ claims a local
  goal to be proven by forward chaining the current facts.  Note that
  $\HENCENAME$ is also equivalent to ``$\FROM{this}~\HAVENAME$''.
  
\item [$\THUSNAME$] abbreviates ``$\THEN~\SHOWNAME$''.  Note that $\THUSNAME$
  is also equivalent to ``$\FROM{this}~\SHOWNAME$''.

\end{descr}

Any goal statement causes some term abbreviations (such as $\Var{thesis}$) to
be bound automatically, see also \S\ref{sec:term-abbrev}.  Furthermore, the
local context of a (non-atomic) goal is provided via the
$rule_context$\indexisarcase{rule-context} case.

\medskip

\begin{warn}
  Isabelle/Isar suffers theory-level goal statements to contain \emph{unbound
    schematic variables}, although this does not conform to the aim of
  human-readable proof documents!  The main problem with schematic goals is
  that the actual outcome is usually hard to predict, depending on the
  behavior of the proof methods applied during the course of reasoning.  Note
  that most semi-automated methods heavily depend on several kinds of implicit
  rule declarations within the current theory context.  As this would also
  result in non-compositional checking of sub-proofs, \emph{local goals} are
  not allowed to be schematic at all.  Nevertheless, schematic goals do have
  their use in Prolog-style interactive synthesis of proven results, usually
  by stepwise refinement via emulation of traditional Isabelle tactic scripts
  (see also \S\ref{sec:tactic-commands}).  In any case, users should know what
  they are doing.
\end{warn}


\subsection{Initial and terminal proof steps}\label{sec:proof-steps}

\indexisarcmd{proof}\indexisarcmd{qed}\indexisarcmd{by}
\indexisarcmd{.}\indexisarcmd{..}\indexisarcmd{sorry}
\begin{matharray}{rcl}
  \isarcmd{proof} & : & \isartrans{proof(prove)}{proof(state)} \\
  \isarcmd{qed} & : & \isartrans{proof(state)}{proof(state) ~|~ theory} \\
  \isarcmd{by} & : & \isartrans{proof(prove)}{proof(state) ~|~ theory} \\
  \isarcmd{.\,.} & : & \isartrans{proof(prove)}{proof(state) ~|~ theory} \\
  \isarcmd{.} & : & \isartrans{proof(prove)}{proof(state) ~|~ theory} \\
  \isarcmd{sorry} & : & \isartrans{proof(prove)}{proof(state) ~|~ theory} \\
\end{matharray}

Arbitrary goal refinement via tactics is considered harmful.  Properly, the
Isar framework admits proof methods to be invoked in two places only.
\begin{enumerate}
\item An \emph{initial} refinement step $\PROOF{m@1}$ reduces a newly stated
  goal to a number of sub-goals that are to be solved later.  Facts are passed
  to $m@1$ for forward chaining, if so indicated by $proof(chain)$ mode.
  
\item A \emph{terminal} conclusion step $\QED{m@2}$ is intended to solve
  remaining goals.  No facts are passed to $m@2$.
\end{enumerate}

The only other (proper) way to affect pending goals in a proof body is by
$\SHOWNAME$, which involves an explicit statement of what is to be solved
eventually.  Thus we avoid the fundamental problem of unstructured tactic
scripts that consist of numerous consecutive goal transformations, with
invisible effects.

\medskip

As a general rule of thumb for good proof style, initial proof methods should
either solve the goal completely, or constitute some well-understood reduction
to new sub-goals.  Arbitrary automatic proof tools that are prone leave a
large number of badly structured sub-goals are no help in continuing the proof
document in an intelligible manner.

Unless given explicitly by the user, the default initial method is ``$rule$'',
which applies a single standard elimination or introduction rule according to
the topmost symbol involved.  There is no separate default terminal method.
Any remaining goals are always solved by assumption in the very last step.

\begin{rail}
  'proof' method?
  ;
  'qed' method?
  ;
  'by' method method?
  ;
  ('.' | '..' | 'sorry')
  ;
\end{rail}

\begin{descr}
  
\item [$\PROOF{m@1}$] refines the goal by proof method $m@1$; facts for
  forward chaining are passed if so indicated by $proof(chain)$ mode.
  
\item [$\QED{m@2}$] refines any remaining goals by proof method $m@2$ and
  concludes the sub-proof by assumption.  If the goal had been $\SHOWNAME$ (or
  $\THUSNAME$), some pending sub-goal is solved as well by the rule resulting
  from the result \emph{exported} into the enclosing goal context.  Thus
  $\QEDNAME$ may fail for two reasons: either $m@2$ fails, or the resulting
  rule does not fit to any pending goal\footnote{This includes any additional
    ``strong'' assumptions as introduced by $\ASSUMENAME$.} of the enclosing
  context.  Debugging such a situation might involve temporarily changing
  $\SHOWNAME$ into $\HAVENAME$, or weakening the local context by replacing
  occurrences of $\ASSUMENAME$ by $\PRESUMENAME$.
  
\item [$\BYY{m@1}{m@2}$] is a \emph{terminal proof}\index{proof!terminal}; it
  abbreviates $\PROOF{m@1}~\QED{m@2}$, but with backtracking across both
  methods.  Debugging an unsuccessful $\BYY{m@1}{m@2}$ commands might be done
  by expanding its definition; in many cases $\PROOF{m@1}$ (or even
  $\APPLY{m@1}$) is already sufficient to see the problem.

\item [``$\DDOT$''] is a \emph{default proof}\index{proof!default}; it
  abbreviates $\BY{rule}$.

\item [``$\DOT$''] is a \emph{trivial proof}\index{proof!trivial}; it
  abbreviates $\BY{this}$.
  
\item [$\SORRY$] is a \emph{fake proof}\index{proof!fake} pretending to solve
  the pending claim without further ado.  This only works in interactive
  development, or if the \texttt{quick_and_dirty} flag is enabled.  Facts
  emerging from fake proofs are not the real thing.  Internally, each theorem
  container is tainted by an oracle invocation, which is indicated as
  ``$[!]$'' in the printed result.
  
  The most important application of $\SORRY$ is to support experimentation and
  top-down proof development.
\end{descr}


\subsection{Fundamental methods and attributes}\label{sec:pure-meth-att}

The following proof methods and attributes refer to basic logical operations
of Isar.  Further methods and attributes are provided by several generic and
object-logic specific tools and packages (see chapters \ref{ch:gen-tools} and
\ref{ch:logics}).

\indexisarmeth{$-$}\indexisarmeth{assumption}
\indexisarmeth{this}\indexisarmeth{rule}\indexisarmeth{rules}
\indexisarattof{Pure}{intro}\indexisarattof{Pure}{elim}
\indexisarattof{Pure}{dest}\indexisarattof{Pure}{rule}
\indexisaratt{OF}\indexisaratt{of}\indexisaratt{where}
\begin{matharray}{rcl}
  - & : & \isarmeth \\
  assumption & : & \isarmeth \\
  this & : & \isarmeth \\
  rule & : & \isarmeth \\
  rules & : & \isarmeth \\[0.5ex]
  intro & : & \isaratt \\
  elim & : & \isaratt \\
  dest & : & \isaratt \\
  rule & : & \isaratt \\[0.5ex]
  OF & : & \isaratt \\
  of & : & \isaratt \\
  where & : & \isaratt \\
\end{matharray}

\begin{rail}
  'rule' thmrefs?
  ;
  'rules' ('!' ?) (rulemod *)
  ;
  rulemod: ('intro' | 'elim' | 'dest') ((('!' | () | '?') nat?) | 'del') ':' thmrefs
  ;
  ('intro' | 'elim' | 'dest') ('!' | () | '?') nat?
  ;
  'rule' 'del'
  ;
  'OF' thmrefs
  ;
  'of' insts ('concl' ':' insts)?
  ;
  'where' ((name | var | typefree | typevar) '=' (type | term) * 'and')
  ;
\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; thus a plain
  \emph{do-nothing} proof step would be ``$\PROOF{-}$'' rather than
  $\PROOFNAME$ alone.
  
\item [$assumption$] solves some goal by a single assumption step.  All given
  facts are guaranteed to participate in the refinement; this means there may
  be only $0$ or $1$ in the first place.  Recall that $\QEDNAME$ (see
  \S\ref{sec:proof-steps}) already concludes any remaining sub-goals by
  assumption, so structured proofs usually need not quote the $assumption$
  method at all.
  
\item [$this$] applies all of the current facts directly as rules.  Recall
  that ``$\DOT$'' (dot) abbreviates ``$\BY{this}$''.
  
\item [$rule~\vec a$] 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 introduction, while with facts it becomes
  elimination.
  
  When no arguments are given, the $rule$ method tries to pick appropriate
  rules automatically, as declared in the current context using the $intro$,
  $elim$, $dest$ attributes (see below).  This is the default behavior of
  $\PROOFNAME$ and ``$\DDOT$'' (double-dot) steps (see
  \S\ref{sec:proof-steps}).
  
\item [$rules$] performs intuitionistic proof search, depending on
  specifically declared rules from the context, or given as explicit
  arguments.  Chained facts are inserted into the goal before commencing proof
  search; ``$rules!$'' means to include the current $prems$ as well.
  
  Rules need to be classified as $intro$, $elim$, or $dest$; here the ``$!$''
  indicator refers to ``safe'' rules, which may be applied aggressively
  (without considering back-tracking later).  Rules declared with ``$?$'' are
  ignored in proof search (the single-step $rule$ method still observes
  these).  An explicit weight annotation may be given as well; otherwise the
  number of rule premises will be taken into account here.
  
\item [$intro$, $elim$, and $dest$] declare introduction, elimination, and
  destruct rules, to be used with the $rule$ and $rules$ methods.  Note that
  the latter will ignore rules declared with ``$?$'', while ``$!$'' are used
  most aggressively.
  
  The classical reasoner (see \S\ref{sec:classical}) introduces its own
  variants of these attributes; use qualified names to access the present
  versions of Isabelle/Pure, i.e.\ $Pure{\dtt}intro$ or $CPure{\dtt}intro$.
  
\item [$rule~del$] undeclares introduction, elimination, or destruct rules.
  
\item [$OF~\vec a$] applies some theorem to given rules $\vec a$ (in
  parallel).  This corresponds to the \texttt{MRS} operator in ML
  \cite[\S5]{isabelle-ref}, but note the reversed order.  Positions may be
  effectively skipped by including ``$\_$'' (underscore) as argument.
  
\item [$of~\vec t$] performs positional instantiation of term variables.  The
  terms $\vec t$ 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 [$where~\vec x = \vec t$] performs named instantiation of schematic type
  and term variables occurring in a theorem.  Schematic variables have to be
  specified on the left-hand side (e.g.\ $?x1\!.\!3$).  The question mark may
  be omitted if the variable name is a plain identifier without index.  As
  type instantiations are inferred from term instantiations, explicit type
  instantiations are seldom necessary.

\end{descr}


\subsection{Term abbreviations}\label{sec:term-abbrev}

\indexisarcmd{let}
\begin{matharray}{rcl}
  \isarcmd{let} & : & \isartrans{proof(state)}{proof(state)} \\
  \isarkeyword{is} & : & syntax \\
\end{matharray}

Abbreviations may be either bound by explicit $\LET{p \equiv t}$ statements,
or by annotating assumptions or goal statements with a list of patterns
``$\ISS{p@1\;\dots}{p@n}$''.  In both cases, higher-order matching is invoked
to bind extra-logical term variables, which may be either named schematic
variables of the form $\Var{x}$, or nameless dummies ``\texttt{_}''
(underscore).\indexisarvar{_@\texttt{_}} Note that in the $\LETNAME$ form the
patterns occur on the left-hand side, while the $\ISNAME$ patterns are in
postfix position.

Polymorphism of term bindings is handled in Hindley-Milner style, similar to
ML.  Type variables referring to local assumptions or open goal statements are
\emph{fixed}, while those of finished results or bound by $\LETNAME$ may occur
in \emph{arbitrary} instances later.  Even though actual polymorphism should
be rarely used in practice, this mechanism is essential to achieve proper
incremental type-inference, as the user proceeds to build up the Isar proof
text from left to right.

\medskip

Term abbreviations are quite different from local definitions as introduced
via $\DEFNAME$ (see \S\ref{sec:proof-context}).  The latter are visible within
the logic as actual equations, while abbreviations disappear during the input
process just after type checking.  Also note that $\DEFNAME$ does not support
polymorphism.

\begin{rail}
  'let' ((term + 'and') '=' term + 'and')
  ;  
\end{rail}

The syntax of $\ISNAME$ patterns follows \railnonterm{termpat} or
\railnonterm{proppat} (see \S\ref{sec:term-decls}).

\begin{descr}
\item [$\LET{\vec p = \vec t}$] binds any text variables in patters $\vec p$
  by simultaneous higher-order matching against terms $\vec t$.
\item [$\IS{\vec p}$] resembles $\LETNAME$, but matches $\vec p$ against the
  preceding statement.  Also note that $\ISNAME$ is not a separate command,
  but part of others (such as $\ASSUMENAME$, $\HAVENAME$ etc.).
\end{descr}

Some \emph{automatic} term abbreviations\index{term abbreviations} for goals
and facts are available as well.  For any open goal,
$\Var{thesis}$\indexisarvar{thesis} refers to its object-level statement,
abstracted over any meta-level parameters (if present).  Likewise,
$\Var{this}$\indexisarvar{this} is bound for fact statements resulting from
assumptions or finished goals.  In case $\Var{this}$ refers to an object-logic
statement that is an application $f(t)$, then $t$ is bound to the special text
variable ``$\dots$''\indexisarvar{\dots} (three dots).  The canonical
application of the latter are calculational proofs (see
\S\ref{sec:calculation}).


\subsection{Block structure}

\indexisarcmd{next}\indexisarcmd{\{}\indexisarcmd{\}}
\begin{matharray}{rcl}
  \NEXT & : & \isartrans{proof(state)}{proof(state)} \\
  \BG & : & \isartrans{proof(state)}{proof(state)} \\
  \EN & : & \isartrans{proof(state)}{proof(state)} \\
\end{matharray}

While Isar is inherently block-structured, opening and closing blocks is
mostly handled rather casually, with little explicit user-intervention.  Any
local goal statement automatically opens \emph{two} blocks, which are closed
again when concluding the sub-proof (by $\QEDNAME$ etc.).  Sections of
different context within a sub-proof may be switched via $\NEXT$, which is
just a single block-close followed by block-open again.  The effect of $\NEXT$
is to reset the local proof context; there is no goal focus involved here!

For slightly more advanced applications, there are explicit block parentheses
as well.  These typically achieve a stronger forward style of reasoning.

\begin{descr}
\item [$\NEXT$] switches to a fresh block within a sub-proof, resetting the
  local context to the initial one.
\item [$\BG$ and $\EN$] explicitly open and close blocks.  Any current facts
  pass through ``$\BG$'' unchanged, while ``$\EN$'' causes any result to be
  \emph{exported} into the enclosing context.  Thus fixed variables are
  generalized, assumptions discharged, and local definitions unfolded (cf.\ 
  \S\ref{sec:proof-context}).  There is no difference of $\ASSUMENAME$ and
  $\PRESUMENAME$ in this mode of forward reasoning --- in contrast to plain
  backward reasoning with the result exported at $\SHOWNAME$ time.
\end{descr}


\subsection{Emulating tactic scripts}\label{sec:tactic-commands}

The Isar provides separate commands to accommodate tactic-style proof scripts
within the same system.  While being outside the orthodox Isar proof language,
these might come in handy for interactive exploration and debugging, or even
actual tactical proof within new-style theories (to benefit from document
preparation, for example).  See also \S\ref{sec:tactics} for actual tactics,
that have been encapsulated as proof methods.  Proper proof methods may be
used in scripts, too.

\indexisarcmd{apply}\indexisarcmd{apply-end}\indexisarcmd{done}
\indexisarcmd{defer}\indexisarcmd{prefer}\indexisarcmd{back}
\indexisarcmd{declare}
\begin{matharray}{rcl}
  \isarcmd{apply}^* & : & \isartrans{proof(prove)}{proof(prove)} \\
  \isarcmd{apply_end}^* & : & \isartrans{proof(state)}{proof(state)} \\
  \isarcmd{done}^* & : & \isartrans{proof(prove)}{proof(state)} \\
  \isarcmd{defer}^* & : & \isartrans{proof}{proof} \\
  \isarcmd{prefer}^* & : & \isartrans{proof}{proof} \\
  \isarcmd{back}^* & : & \isartrans{proof}{proof} \\
  \isarcmd{declare}^* & : & \isartrans{theory}{theory} \\
\end{matharray}

\railalias{applyend}{apply\_end}
\railterm{applyend}

\begin{rail}
  ( 'apply' | applyend ) method
  ;
  'defer' nat?
  ;
  'prefer' nat
  ;
  'declare' locale? (thmrefs + 'and')
  ;
\end{rail}

\begin{descr}

\item [$\APPLY{m}$] applies proof method $m$ in initial position, but unlike
  $\PROOFNAME$ it retains ``$proof(prove)$'' mode.  Thus consecutive method
  applications may be given just as in tactic scripts.
  
  Facts are passed to $m$ as indicated by the goal's forward-chain mode, and
  are \emph{consumed} afterwards.  Thus any further $\APPLYNAME$ command would
  always work in a purely backward manner.
  
\item [$\isarkeyword{apply_end}~(m)$] applies proof method $m$ as if in
  terminal position.  Basically, this simulates a multi-step tactic script for
  $\QEDNAME$, but may be given anywhere within the proof body.
  
  No facts are passed to $m$.  Furthermore, the static context is that of the
  enclosing goal (as for actual $\QEDNAME$).  Thus the proof method may not
  refer to any assumptions introduced in the current body, for example.
  
\item [$\isarkeyword{done}$] completes a proof script, provided that the
  current goal state is solved completely.  Note that actual structured proof
  commands (e.g.\ ``$\DOT$'' or $\SORRY$) may be used to conclude proof
  scripts as well.

\item [$\isarkeyword{defer}~n$ and $\isarkeyword{prefer}~n$] shuffle the list
  of pending goals: $defer$ puts off goal $n$ to the end of the list ($n = 1$
  by default), while $prefer$ brings goal $n$ to the top.
  
\item [$\isarkeyword{back}$] does back-tracking over the result sequence of
  the latest proof command.  Basically, any proof command may return multiple
  results.
  
\item [$\isarkeyword{declare}~thms$] declares theorems to the current theory
  context (or the specified locale, see also \S\ref{sec:locale}).  No theorem
  binding is involved here, unlike $\isarkeyword{theorems}$ or
  $\isarkeyword{lemmas}$ (cf.\ \S\ref{sec:axms-thms}), so
  $\isarkeyword{declare}$ only has the effect of applying attributes as
  included in the theorem specification.

\end{descr}

Any proper Isar proof method may be used with tactic script commands such as
$\APPLYNAME$.  A few additional emulations of actual tactics are provided as
well; these would be never used in actual structured proofs, of course.


\subsection{Meta-linguistic features}

\indexisarcmd{oops}
\begin{matharray}{rcl}
  \isarcmd{oops} & : & \isartrans{proof}{theory} \\
\end{matharray}

The $\OOPS$ command discontinues the current proof attempt, while considering
the partial proof text as properly processed.  This is conceptually quite
different from ``faking'' actual proofs via $\SORRY$ (see
\S\ref{sec:proof-steps}): $\OOPS$ does not observe the proof structure at all,
but goes back right to the theory level.  Furthermore, $\OOPS$ does not
produce any result theorem --- there is no intended claim to be able to
complete the proof anyhow.

A typical application of $\OOPS$ is to explain Isar proofs \emph{within} the
system itself, in conjunction with the document preparation tools of Isabelle
described in \cite{isabelle-sys}.  Thus partial or even wrong proof attempts
can be discussed in a logically sound manner.  Note that the Isabelle {\LaTeX}
macros can be easily adapted to print something like ``$\dots$'' instead of an
``$\OOPS$'' keyword.

\medskip The $\OOPS$ command is undo-able, unlike $\isarkeyword{kill}$ (see
\S\ref{sec:history}).  The effect is to get back to the theory just before the
opening of the proof.


\section{Other commands}

\subsection{Diagnostics}

\indexisarcmd{pr}\indexisarcmd{thm}\indexisarcmd{term}
\indexisarcmd{prop}\indexisarcmd{typ}
\begin{matharray}{rcl}
  \isarcmd{pr}^* & : & \isarkeep{\cdot} \\
  \isarcmd{thm}^* & : & \isarkeep{theory~|~proof} \\
  \isarcmd{term}^* & : & \isarkeep{theory~|~proof} \\
  \isarcmd{prop}^* & : & \isarkeep{theory~|~proof} \\
  \isarcmd{typ}^* & : & \isarkeep{theory~|~proof} \\
  \isarcmd{prf}^* & : & \isarkeep{theory~|~proof} \\
  \isarcmd{full_prf}^* & : & \isarkeep{theory~|~proof} \\
\end{matharray}

These diagnostic commands assist interactive development.  Note that $undo$
does not apply here, the theory or proof configuration is not changed.

\begin{rail}
  'pr' modes? nat? (',' nat)?
  ;
  'thm' modes? thmrefs
  ;
  'term' modes? term
  ;
  'prop' modes? prop
  ;
  'typ' modes? type
  ;
  'prf' modes? thmrefs?
  ;
  'full\_prf' modes? thmrefs?
  ;

  modes: '(' (name + ) ')'
  ;
\end{rail}

\begin{descr}
\item [$\isarkeyword{pr}~goals, prems$] prints the current proof state (if
  present), including the proof context, current facts and goals.  The
  optional limit arguments affect the number of goals and premises to be
  displayed, which is initially 10 for both.  Omitting limit values leaves the
  current setting unchanged.
\item [$\isarkeyword{thm}~\vec a$] retrieves theorems from the current theory
  or proof context.  Note that any attributes included in the theorem
  specifications are applied to a temporary context derived from the current
  theory or proof; the result is discarded, i.e.\ attributes involved in $\vec
  a$ do not have any permanent effect.
\item [$\isarkeyword{term}~t$ and $\isarkeyword{prop}~\phi$] read, type-check
  and print terms or propositions according to the current theory or proof
  context; the inferred type of $t$ is output as well.  Note that these
  commands are also useful in inspecting the current environment of term
  abbreviations.
\item [$\isarkeyword{typ}~\tau$] reads and prints types of the meta-logic
  according to the current theory or proof context.
\item [$\isarkeyword{prf}$] displays the (compact) proof term of the current
  proof state (if present), or of the given theorems. Note that this
  requires proof terms to be switched on for the current object logic
  (see the ``Proof terms'' section of the Isabelle reference manual
  for information on how to do this).
\item [$\isarkeyword{full_prf}$] is like $\isarkeyword{prf}$, but displays
  the full proof term, i.e.\ also displays information omitted in
  the compact proof term, which is denoted by ``$_$'' placeholders there.
\end{descr}

All of the diagnostic commands above admit a list of $modes$ to be specified,
which is appended to the current print mode (see also \cite{isabelle-ref}).
Thus the output behavior may be modified according particular print mode
features.  For example, $\isarkeyword{pr}~(latex~xsymbols~symbols)$ would
print the current proof state with mathematical symbols and special characters
represented in {\LaTeX} source, according to the Isabelle style
\cite{isabelle-sys}.

Note that antiquotations (cf.\ \S\ref{sec:antiq}) provide a more systematic
way to include formal items into the printed text document.


\subsection{Inspecting the context}

\indexisarcmd{print-facts}\indexisarcmd{print-binds}
\indexisarcmd{print-commands}\indexisarcmd{print-syntax}
\indexisarcmd{print-methods}\indexisarcmd{print-attributes}
\indexisarcmd{thms-containing}\indexisarcmd{thm-deps}
\indexisarcmd{print-theorems}
\begin{matharray}{rcl}
  \isarcmd{print_commands}^* & : & \isarkeep{\cdot} \\
  \isarcmd{print_syntax}^* & : & \isarkeep{theory~|~proof} \\
  \isarcmd{print_methods}^* & : & \isarkeep{theory~|~proof} \\
  \isarcmd{print_attributes}^* & : & \isarkeep{theory~|~proof} \\
  \isarcmd{print_theorems}^* & : & \isarkeep{theory~|~proof} \\
  \isarcmd{thms_containing}^* & : & \isarkeep{theory~|~proof} \\
  \isarcmd{thms_deps}^* & : & \isarkeep{theory~|~proof} \\
  \isarcmd{print_facts}^* & : & \isarkeep{proof} \\
  \isarcmd{print_binds}^* & : & \isarkeep{proof} \\
\end{matharray}

\railalias{thmscontaining}{thms\_containing}
\railterm{thmscontaining}

\railalias{thmdeps}{thm\_deps}
\railterm{thmdeps}

\begin{rail}
  thmscontaining (('(' nat ')')?) ((criterion '-'?) * )  
  % no idea why this needs to be backwards (should be '-'? criterion),
  % but it comes out right this way
  ;
  criterion: 'intro' | 'elim' | 'dest' | 'term' | 'name' ':' string
  ;
  thmdeps thmrefs
  ;
\end{rail}

These commands print certain parts of the theory and proof context.  Note that
there are some further ones available, such as for the set of rules declared
for simplifications.

\begin{descr}
  
\item [$\isarkeyword{print_commands}$] prints Isabelle's outer theory syntax,
  including keywords and command.
  
\item [$\isarkeyword{print_syntax}$] prints the inner syntax of types and
  terms, depending on the current context.  The output can be very verbose,
  including grammar tables and syntax translation rules.  See \cite[\S7,
  \S8]{isabelle-ref} for further information on Isabelle's inner syntax.
  
\item [$\isarkeyword{print_methods}$] prints all proof methods available in
  the current theory context.
  
\item [$\isarkeyword{print_attributes}$] prints all attributes available in
  the current theory context.
  
\item [$\isarkeyword{print_theorems}$] prints theorems available in the
  current theory context.
  
  In interactive mode this actually refers to the theorems left by the last
  transaction; this allows to inspect the result of advanced definitional
  packages, such as $\isarkeyword{datatype}$.
  
\item [$\isarkeyword{thms_containing}~\vec c$] retrieves facts from the theory
  or proof context matching all of the search criteria $\vec c$. Valid
  criteria are $\isarkeyword{intro}$, $\isarkeyword{elim}$, and
  $\isarkeyword{dest}$, selecting theorems that match the current goal
  as introduction, elimination or destruction rules respectively,
  $\isarkeyword{name}$ $:$ $s$, selecting all theorems that contain
  $s$ in their fully qualified name, and finally any term $t$ which
  may contain occurrences of ``$\_$'' and type restrictions, selecting
  theorems that contain the pattern $t$. Criteria can be preceded by
  \texttt{-} to select theorems that do \emph{not} match. Note that
  giving the empty list of criteria yields \emph{all} currently known
  facts.  An optional limit for the number of printed facts may be
  given; the default is 40.
  
\item [$\isarkeyword{thm_deps}~\vec a$] visualizes dependencies of facts,
  using Isabelle's graph browser tool (see also \cite{isabelle-sys}).
  
\item [$\isarkeyword{print_facts}$] prints any named facts of the current
  context, including assumptions and local results.
  
\item [$\isarkeyword{print_binds}$] prints all term abbreviations present in
  the context.

\end{descr}


\subsection{History commands}\label{sec:history}

\indexisarcmd{undo}\indexisarcmd{redo}\indexisarcmd{kill}
\begin{matharray}{rcl}
  \isarcmd{undo}^{{*}{*}} & : & \isarkeep{\cdot} \\
  \isarcmd{redo}^{{*}{*}} & : & \isarkeep{\cdot} \\
  \isarcmd{kill}^{{*}{*}} & : & \isarkeep{\cdot} \\
\end{matharray}

The Isabelle/Isar top-level maintains a two-stage history, for theory and
proof state transformation.  Basically, any command can be undone using
$\isarkeyword{undo}$, excluding mere diagnostic elements.  Its effect may be
revoked via $\isarkeyword{redo}$, unless the corresponding
$\isarkeyword{undo}$ step has crossed the beginning of a proof or theory.  The
$\isarkeyword{kill}$ command aborts the current history node altogether,
discontinuing a proof or even the whole theory.  This operation is \emph{not}
undo-able.

\begin{warn}
  History commands should never be used with user interfaces such as
  Proof~General \cite{proofgeneral,Aspinall:TACAS:2000}, which takes care of
  stepping forth and back itself.  Interfering by manual $\isarkeyword{undo}$,
  $\isarkeyword{redo}$, or even $\isarkeyword{kill}$ commands would quickly
  result in utter confusion.
\end{warn}


\subsection{System operations}

\indexisarcmd{cd}\indexisarcmd{pwd}\indexisarcmd{use-thy}\indexisarcmd{use-thy-only}
\indexisarcmd{update-thy}\indexisarcmd{update-thy-only}\indexisarcmd{display-drafts}
\indexisarcmd{print-drafts}
\begin{matharray}{rcl}
  \isarcmd{cd}^* & : & \isarkeep{\cdot} \\
  \isarcmd{pwd}^* & : & \isarkeep{\cdot} \\
  \isarcmd{use_thy}^* & : & \isarkeep{\cdot} \\
  \isarcmd{use_thy_only}^* & : & \isarkeep{\cdot} \\
  \isarcmd{update_thy}^* & : & \isarkeep{\cdot} \\
  \isarcmd{update_thy_only}^* & : & \isarkeep{\cdot} \\
  \isarcmd{display_drafts}^* & : & \isarkeep{\cdot} \\
  \isarcmd{print_drafts}^* & : & \isarkeep{\cdot} \\
\end{matharray}

\railalias{usethy}{use\_thy}
\railterm{usethy}
\railalias{usethyonly}{use\_thy\_only}
\railterm{usethyonly}
\railalias{updatethy}{update\_thy}
\railterm{updatethy}
\railalias{updatethyonly}{update\_thy\_only}
\railterm{updatethyonly}
\railalias{displaydrafts}{display\_drafts}
\railterm{displaydrafts}
\railalias{printdrafts}{print\_drafts}
\railterm{printdrafts}

\begin{rail}
  ('cd' | usethy | usethyonly | updatethy | updatethyonly) name
  ;
  (displaydrafts | printdrafts) (name +)
  ;
\end{rail}

\begin{descr}
\item [$\isarkeyword{cd}~path$] changes the current directory of the Isabelle
  process.
\item [$\isarkeyword{pwd}~$] prints the current working directory.
\item [$\isarkeyword{use_thy}$, $\isarkeyword{use_thy_only}$,
  $\isarkeyword{update_thy}$, $\isarkeyword{update_thy_only}$] load some
  theory given as $name$ argument.  These commands are basically the same as
  the corresponding ML functions\footnote{The ML versions also change the
    implicit theory context to that of the theory loaded.}  (see also
  \cite[\S1,\S6]{isabelle-ref}).  Note that both the ML and Isar versions may
  load new- and old-style theories alike.
\item [$\isarkeyword{display_drafts}~paths$ and
  $\isarkeyword{print_drafts}~paths$] perform simple output of a given list of
  raw source files.  Only those symbols that do not require additional
  {\LaTeX} packages are displayed properly, everything else is left verbatim.
\end{descr}

These system commands are scarcely used when working with the Proof~General
interface, since loading of theories is done transparently.

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