\chapter{Basic Isar Language Elements}\label{ch:pure-syntax}
Subsequently, we introduce the main part of Pure Isar 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 by most object logics. Chapter~\ref{ch:hol-tools}
refers to actual object-logic specific elements of Isabelle/HOL.
\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 outcome. Typical
examples are diagnostic commands that print terms or theorems according to the
current context; other commands even emulate old-style tactical theorem
proving.
\section{Theory commands}
\subsection{Defining theories}\label{sec:begin-thy}
\indexisarcmd{header}\indexisarcmd{theory}\indexisarcmd{end}\indexisarcmd{context}
\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 actual command of any theory has to be $\THEORY$, starting 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 theory
context may be also changed by $\CONTEXT$ without creating a new theory. In
both cases, $\END$ concludes the theory development; it has to be the very
last command of any theory file.
\begin{rail}
'header' text
;
'theory' name '=' (name + '+') filespecs? ':'
;
'context' name
;
'end'
;;
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$] commences a new theory $A$
based on existing ones $B@1 + \cdots + B@n$. Isabelle's theory loader
system ensures that any of the base theories are properly loaded (and fully
up-to-date when $\THEORY$ is executed interactively). The optional
$\isarkeyword{files}$ specification declares additional dependencies on ML
files. Unless put in parentheses, any file will be loaded immediately via
$\isarcmd{use}$ (see also \S\ref{sec:ML}). The optional ML file
\texttt{$A$.ML} that may be associated with any theory should \emph{not} be
included in $\isarkeyword{files}$, 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.
\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{Theory 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.\footnote{The latter feature is not yet supported.
Nevertheless, any source text of the form
``\texttt{\at\ttlbrace$\dots$\ttrbrace}'' should be considered as reserved
for future use.}
\item [$\isarkeyword{text_raw}$] inserts {\LaTeX} source into the output,
without additional markup. Thus the full range of document manipulations
becomes available. A typical application would be to emit
\verb,\begin{comment}, and \verb,\end{comment}, commands to exclude certain
parts from the final document.\footnote{This requires the \texttt{comment}
package to be included in {\LaTeX}, of course.}
\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}{rcl}
\isarcmd{classes} & : & \isartrans{theory}{theory} \\
\isarcmd{classrel} & : & \isartrans{theory}{theory} \\
\isarcmd{defaultsort} & : & \isartrans{theory}{theory} \\
\end{matharray}
\begin{rail}
'classes' (classdecl comment? +)
;
'classrel' nameref '<' nameref comment?
;
'defaultsort' sort comment?
;
\end{rail}
\begin{descr}
\item [$\isarkeyword{classes}~c<\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<c@2$] states a subclass relation between
existing classes $c@1$ and $c@2$. This is done axiomatically! The
$\isarkeyword{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 new object-logics.
\end{descr}
\subsection{Primitive types and type abbreviations}\label{sec:types-pure}
\indexisarcmd{typedecl}\indexisarcmd{types}\indexisarcmd{nonterminals}\indexisarcmd{arities}
\begin{matharray}{rcl}
\isarcmd{types} & : & \isartrans{theory}{theory} \\
\isarcmd{typedecl} & : & \isartrans{theory}{theory} \\
\isarcmd{nonterminals} & : & \isartrans{theory}{theory} \\
\isarcmd{arities} & : & \isartrans{theory}{theory} \\
\end{matharray}
\begin{rail}
'types' (typespec '=' type infix? comment? +)
;
'typedecl' typespec infix? comment?
;
'nonterminals' (name +) comment?
;
'arities' (nameref '::' arity comment? +)
;
\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 object-logics such as
Isabelle/HOL override $\isarkeyword{typedecl}$ by their own version.
\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 $\isarkeyword{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' (constdecl +)
;
'defs' (axmdecl prop comment? +)
;
'constdefs' (constdecl prop comment? +)
;
constdecl: name '::' type mixfix? comment?
;
\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.
\item [$\isarkeyword{constdefs}~c::\sigma~eqn$] combines declarations and
definitions of constants, using the canonical name $c_def$ for the
definitional axiom.
\end{descr}
\subsection{Syntax and translations}\label{sec:syn-trans}
\indexisarcmd{syntax}\indexisarcmd{translations}
\begin{matharray}{rcl}
\isarcmd{syntax} & : & \isartrans{theory}{theory} \\
\isarcmd{translations} & : & \isartrans{theory}{theory} \\
\end{matharray}
\begin{rail}
'syntax' ('(' name 'output'? ')')? (constdecl +)
;
'translations' (transpat ('==' | '=>' | '<=') transpat comment? +)
;
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 \texttt{output}
flag is given, all productions are added both to the input and output
grammar.
\item [$\isarkeyword{translations}~rules$] specifies syntactic translation
rules (i.e.\ \emph{macros}): parse~/ print rules (\texttt{==}), parse rules
(\texttt{=>}), or print rules (\texttt{<=}). Translation patterns may be
prefixed by the syntactic category to be used for parsing; the default is
\texttt{logic}.
\end{descr}
\subsection{Axioms and theorems}
\indexisarcmd{axioms}\indexisarcmd{theorems}\indexisarcmd{lemmas}
\begin{matharray}{rcl}
\isarcmd{axioms} & : & \isartrans{theory}{theory} \\
\isarcmd{theorems} & : & \isartrans{theory}{theory} \\
\isarcmd{lemmas} & : & \isartrans{theory}{theory} \\
\end{matharray}
\begin{rail}
'axioms' (axmdecl prop comment? +)
;
('theorems' | 'lemmas') (thmdef? thmrefs comment? + '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
actual proven theorems.
\item [$\isarkeyword{theorems}~a = \vec b$] stores lists of existing theorems.
Typical applications would also involve attributes, to declare Simplifier
rules, for example.
\item [$\isarkeyword{lemmas}$] is similar to $\isarkeyword{theorems}$, but
tags the results as ``lemma''.
\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}
'global' comment?
;
'local' comment?
;
'hide' name (nameref + ) comment?
;
\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 deliberately.
\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 comment?
;
\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
$\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)))
\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}
Syntax translation functions written in ML admit almost arbitrary
manipulations of Isabelle's inner syntax. Any of the above commands have a
single \railqtoken{text} argument that refers to an ML expression of
appropriate type.
\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}
See \cite[\S8]{isabelle-ref} for more information on syntax transformations.
\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 comment?
;
\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 an internal 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}
\subsection{Proof 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{Proof context}\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]$.
\begin{rail}
'fix' (vars + 'and') comment?
;
('assume' | 'presume') (assm comment? + 'and')
;
'def' thmdecl? \\ var '==' term termpat? comment?
;
var: name ('::' type)?
;
vars: (name+) ('::' type)?
;
assm: thmdecl? (prop proppat? +)
;
\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}
\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)} \\
\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). Note that the special theorem name
$this$\indexisarthm{this} refers to the most recently established facts.
\begin{rail}
'note' (thmdef? thmrefs comment? + 'and')
;
'then' comment?
;
('from' | 'with') (thmrefs comment? + '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'' (cf.\
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~facts}$; thus the forward
chaining is from earlier facts together with the current ones.
\end{descr}
Basic proof methods (such as $rule$, see \S\ref{sec:pure-meth-att}) 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{_}}
Forward chaining with an empty list of theorems is the same as not chaining.
Thus $\FROM{nothing}$ has no effect apart from entering $prove(chain)$ mode,
since $nothing$\indexisarthm{nothing} is bound to the empty list of facts.
\subsection{Goal statements}
\indexisarcmd{theorem}\indexisarcmd{lemma}
\indexisarcmd{have}\indexisarcmd{show}\indexisarcmd{hence}\indexisarcmd{thus}
\begin{matharray}{rcl}
\isarcmd{theorem} & : & \isartrans{theory}{proof(prove)} \\
\isarcmd{lemma} & : & \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}
Proof mode is entered from theory mode by initial goal commands $\THEOREMNAME$
and $\LEMMANAME$. New local goals may be claimed within proof mode as well.
Four variants are available, indicating whether the result is meant to solve
some pending goal or whether forward chaining is indicated.
\begin{rail}
('theorem' | 'lemma') goal
;
('have' | 'show' | 'hence' | 'thus') goal
;
goal: thmdecl? prop proppat? comment?
;
\end{rail}
\begin{descr}
\item [$\THEOREM{a}{\phi}$] enters proof mode with $\phi$ as main goal,
eventually resulting in some theorem $\turn \phi$ to be put back into the
theory.
\item [$\LEMMA{a}{\phi}$] is similar to $\THEOREMNAME$, but tags the result as
``lemma''.
\item [$\HAVE{a}{\phi}$] claims a local goal, eventually resulting in a
theorem with the current assumption context as hypotheses.
\item [$\SHOW{a}{\phi}$] is similar to $\HAVE{a}{\phi}$, but solves some
pending goal with the result \emph{exported} into the corresponding context
(cf.\ \S\ref{sec:proof-context}).
\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}
Note that any goal statement causes some term abbreviations (such as
$\Var{thesis}$, $\dots$) to be bound automatically, see also
\S\ref{sec:term-abbrev}. Furthermore, the local context of a (non-atomic)
goal is provided via the case name $antecedent$\indexisarcase{antecedent}, see
also \S\ref{sec:cases}.
\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 is by $\SHOWNAME$, which
involves an explicit statement of what is to be solved.
\medskip
Also note that 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 any intelligible
way.
\medskip
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' interest? meth? comment?
;
'qed' meth? comment?
;
'by' meth meth? comment?
;
('.' | '..' | 'sorry') comment?
;
meth: method interest?
;
\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
some 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}$, with backtracking across both methods,
though. Debugging an unsuccessful $\BYY{m@1}{m@2}$ commands might be done
by expanding its definition; in many cases $\PROOF{m@1}$ is already
sufficient to see what is going wrong.
\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}; provided that the
\texttt{quick_and_dirty} flag is enabled, $\SORRY$ pretends to solve the
goal without further ado. Of course, the result would be a fake theorem
only, involving some oracle in its internal derivation object (this is
indicated as ``$[!]$'' in the printed result). The main 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:hol-tools}).
\indexisarmeth{assumption}\indexisarmeth{this}\indexisarmeth{rule}\indexisarmeth{$-$}
\indexisaratt{intro}\indexisaratt{elim}\indexisaratt{dest}
\indexisaratt{OF}\indexisaratt{of}
\begin{matharray}{rcl}
assumption & : & \isarmeth \\
this & : & \isarmeth \\
rule & : & \isarmeth \\
- & : & \isarmeth \\
OF & : & \isaratt \\
of & : & \isaratt \\
intro & : & \isaratt \\
elim & : & \isaratt \\
dest & : & \isaratt \\
delrule & : & \isaratt \\
\end{matharray}
\begin{rail}
'rule' thmrefs?
;
'OF' thmrefs
;
'of' insts ('concl' ':' insts)?
;
\end{rail}
\begin{descr}
\item [$assumption$] solves some goal by a single assumption step. Any facts
given (${} \le 1$) are guaranteed to participate in the refinement. Recall
that $\QEDNAME$ (see \S\ref{sec:proof-steps}) already concludes any
remaining sub-goals by assumption.
\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 \emph{introduction}, while with facts it
becomes \emph{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 [``$-$''] 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 [$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
skipped by including ``$\_$'' (underscore) as argument.
\item [$of~\vec t$] performs positional instantiation. 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 [$intro$, $elim$, and $dest$] declare introduction, elimination, and
destruct rules, respectively. Note that the classical reasoner (see
\S\ref{sec:classical-basic}) introduces different versions of these
attributes, and the $rule$ method, too. In object-logics with classical
reasoning enabled, the latter version should be used all the time to avoid
confusion!
\item [$delrule$] undeclares introduction or elimination rules.
\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, as in 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.
\medskip
Term abbreviations are quite different from actual 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 comment? + 'and')
;
\end{rail}
The syntax of $\ISNAME$ patterns follows \railnonterm{termpat} or
\railnonterm{proppat} (see \S\ref{sec:term-pats}).
\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}
A few \emph{automatic} term abbreviations\index{term abbreviations} for goals
and facts are available as well. For any open goal,
$\Var{thesis_prop}$\indexisarvar{thesis-prop} refers to the full proposition
(which may be a rule), $\Var{thesis_concl}$\indexisarvar{thesis-concl} to its
(atomic) conclusion, and $\Var{thesis}$\indexisarvar{thesis} to its
object-level statement. The latter two abstract over any meta-level
parameters.
Fact statements resulting from assumptions or finished goals are bound as
$\Var{this_prop}$\indexisarvar{this-prop},
$\Var{this_concl}$\indexisarvar{this-concl}, and
$\Var{this}$\indexisarvar{this}, similar to $\Var{thesis}$ above. 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}
\railalias{lbrace}{\ttlbrace}
\railterm{lbrace}
\railalias{rbrace}{\ttrbrace}
\railterm{rbrace}
\begin{rail}
'next' comment?
;
lbrace comment?
;
rbrace comment?
;
\end{rail}
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. Thus the effect of
$\NEXT$ 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:tactical-proof}
The following elements emulate unstructured tactic scripts to some extent.
While these are anathema for writing proper Isar proof documents, they 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).
\indexisarcmd{apply}\indexisarcmd{done}\indexisarcmd{apply-end}
\indexisarcmd{defer}\indexisarcmd{prefer}\indexisarcmd{back}
\begin{matharray}{rcl}
\isarcmd{apply}^* & : & \isartrans{proof(prove)}{proof(prove)} \\
\isarcmd{done}^* & : & \isartrans{proof(prove)}{proof(state)} \\
\isarcmd{apply_end}^* & : & \isartrans{proof(state)}{proof(state)} \\
\isarcmd{defer}^* & : & \isartrans{proof}{proof} \\
\isarcmd{prefer}^* & : & \isartrans{proof}{proof} \\
\isarcmd{back}^* & : & \isartrans{proof}{proof} \\
\end{matharray}
\railalias{applyend}{apply\_end}
\railterm{applyend}
\begin{rail}
'apply' method comment?
;
'done' comment?
;
applyend method comment?
;
'defer' nat? comment?
;
'prefer' nat comment?
;
\end{rail}
\begin{descr}
\item [$\isarkeyword{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 $\isarkeyword{apply}$
command would always work in a purely backward manner.
\item [$\isarkeyword{done}$] completes a proof script, provided that the
current goal state is already solved completely. Note that actual
structured proof commands (e.g.\ ``$\DOT$'' or $\SORRY$) may be used to
conclude proof scripts as well.
\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{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.\footnote{Unlike the ML function \texttt{back}
\cite{isabelle-ref}, the Isar command does not search upwards for further
branch points.} Basically, any proof command may return multiple results.
\end{descr}
Any proper Isar proof method may be used with tactic script commands such as
$\isarkeyword{apply}$. A few additional emulations of actual tactics are
provided as well; these would be never used in actual structured proofs, of
course.
\indexisarmeth{tactic}\indexisarmeth{insert}
\indexisarmeth{res-inst-tac}\indexisarmeth{eres-inst-tac}
\indexisarmeth{dres-inst-tac}\indexisarmeth{forw-inst-tac}
\indexisarmeth{subgoal-tac}
\begin{matharray}{rcl}
tactic^* & : & \isarmeth \\
insert^* & : & \isarmeth \\
res_inst_tac^* & : & \isarmeth \\
eres_inst_tac^* & : & \isarmeth \\
dres_inst_tac^* & : & \isarmeth \\
forw_inst_tac^* & : & \isarmeth \\
subgoal_tac^* & : & \isarmeth \\
\end{matharray}
\railalias{resinsttac}{res\_inst\_tac}
\railterm{resinsttac}
\railalias{eresinsttac}{eres\_inst\_tac}
\railterm{eresinsttac}
\railalias{dresinsttac}{dres\_inst\_tac}
\railterm{dresinsttac}
\railalias{forwinsttac}{forw\_inst\_tac}
\railterm{forwinsttac}
\railalias{subgoaltac}{subgoal\_tac}
\railterm{subgoaltac}
\begin{rail}
'tactic' text
;
'insert' thmrefs
;
( resinsttac | eresinsttac | dresinsttac | forwinsttac ) goalspec? \\
((name '=' term) + 'and') 'in' thmref
;
subgoaltac goalspec? prop
;
\end{rail}
\begin{descr}
\item [$tactic~text$] produces a proof method from any ML text of type
\texttt{tactic}. Apart from the usual ML environment and the current
implicit theory context, the ML code may refer to the following locally
bound values:
%%FIXME ttbox produces too much trailing space (why?)
{\footnotesize\begin{verbatim}
val ctxt : Proof.context
val facts : thm list
val thm : string -> thm
val thms : string -> thm list
\end{verbatim}}
Here \texttt{ctxt} refers to the current proof context, \texttt{facts}
indicates any current facts for forward-chaining, and
\texttt{thm}~/~\texttt{thms} retrieve named facts (including global
theorems) from the context.
\item [$insert~\vec a$] inserts theorems as facts into all goals of the proof
state; the current facts indicated for forward chaining are ignored!
\item [$res_inst_tac$ etc.] do resolution of rules with explicit
instantiation. This works the same way as the corresponding ML tactics, see
\cite[\S3]{isabelle-ref}.
It is very important to note that the instantiations are read and
type-checked according to the dynamic goal state, rather than the static
proof context! In particular, locally fixed variables and term
abbreviations may not be included in the term specifications.
\item [$subgoal_tac~\phi$] emulates the ML tactic of the same name, see
\cite[\S3]{isabelle-ref}. Syntactically, the given proposition is handled
as the instantiations in $res_inst_tac$ etc.
Note that the proper Isar command $\PRESUMENAME$ achieves a similar effect
as $subgoal_tac$.
\end{descr}
\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 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 undoable, unlike $\isarkeyword{kill}$ (see
\S\ref{sec:history}). The effect is to get back to the theory \emph{before}
the opening of the proof.
\section{Other commands}
\subsection{Diagnostics}\label{sec:diag}
\indexisarcmd{pr}\indexisarcmd{thm}\indexisarcmd{term}\indexisarcmd{prop}\indexisarcmd{typ}
\indexisarcmd{print-facts}\indexisarcmd{print-binds}
\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{print_facts}^* & : & \isarkeep{proof} \\
\isarcmd{print_binds}^* & : & \isarkeep{proof} \\
\end{matharray}
These commands are not part of the actual Isabelle/Isar syntax, but assist
interactive development. Also note that $undo$ does not apply here, since the
theory or proof configuration is not changed.
\begin{rail}
'pr' modes? nat?
;
'thm' modes? thmrefs
;
'term' modes? term
;
'prop' modes? prop
;
'typ' modes? type
;
modes: '(' (name + ) ')'
;
\end{rail}
\begin{descr}
\item [$\isarkeyword{pr}~n$] prints the current proof state (if present),
including the proof context, current facts and goals. The optional argument
$n$ affects the implicit limit of goals to be displayed, which is initially
10. Omitting the limit leaves the current value 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$, $\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{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}
The basic 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}. The resulting text can be directly pasted into a
\verb,\begin{isabelle},\dots\verb,\end{isabelle}, environment. Note that
$\isarkeyword{pr}~(latex)$ is sufficient to achieve the same output, if the
current Isabelle session has the other modes already activated, say due to
some particular user interface configuration such as Proof~General
\cite{proofgeneral,Aspinall:TACAS:2000} with X-Symbol mode \cite{x-symbol}.
\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 the
$\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}
undoable.
\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}
\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} \\
\end{matharray}
\begin{descr}
\item [$\isarkeyword{cd}~name$] 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.
\end{descr}
These system commands are scarcely used when working with the Proof~General
interface, since loading of theories is done fully transparently.
%%% Local Variables:
%%% mode: latex
%%% TeX-master: "isar-ref"
%%% End: