\chapter{Generic tools and packages}\label{ch:gen-tools}
\section{Specification commands}
\subsection{Derived specifications}
\indexisarcmd{axiomatization}
\indexisarcmd{definition}\indexisaratt{defn}
\indexisarcmd{abbreviation}\indexisarcmd{print-abbrevs}
\indexisarcmd{notation}\indexisarcmd{no-notation}
\begin{matharray}{rcll}
\isarcmd{axiomatization} & : & \isarkeep{local{\dsh}theory} & (axiomatic!)\\
\isarcmd{definition} & : & \isarkeep{local{\dsh}theory} \\
defn & : & \isaratt \\
\isarcmd{abbreviation} & : & \isarkeep{local{\dsh}theory} \\
\isarcmd{print_abbrevs}^* & : & \isarkeep{theory~|~proof} \\
\isarcmd{notation} & : & \isarkeep{local{\dsh}theory} \\
\isarcmd{no_notation} & : & \isarkeep{local{\dsh}theory} \\
\end{matharray}
These specification mechanisms provide a slightly more abstract view
than the underlying primitives of $\CONSTS$, $\DEFS$ (see
\S\ref{sec:consts}), and $\isarkeyword{axioms}$ (see
\S\ref{sec:axms-thms}). In particular, type-inference is commonly
available, and result names need not be given.
\begin{rail}
'axiomatization' target? fixes? ('where' specs)?
;
'definition' target? (decl 'where')? thmdecl? prop
;
'abbreviation' target? mode? (decl 'where')? prop
;
('notation' | 'no\_notation') target? mode? (nameref structmixfix + 'and')
;
fixes: ((name ('::' type)? mixfix? | vars) + 'and')
;
specs: (thmdecl? props + 'and')
;
decl: name ('::' type)? mixfix?
;
\end{rail}
\begin{descr}
\item $\isarkeyword{axiomatization} ~ c@1 \dots c@n ~
\isarkeyword{where} ~ A@1 \dots A@m$ introduces several constants
simultaneously and states axiomatic properties for these. The
constants are marked as being specified once and for all, which
prevents additional specifications being issued later on.
Note that axiomatic specifications are only appropriate when
declaring a new logical system. Normal applications should only use
definitional mechanisms!
\item $\isarkeyword{definition}~c~\isarkeyword{where}~eq$ produces an
internal definition $c \equiv t$ according to the specification
given as $eq$, which is then turned into a proven fact. The given
proposition may deviate from internal meta-level equality according
to the rewrite rules declared as $defn$ by the object-logic. This
typically covers object-level equality $x = t$ and equivalence $A
\leftrightarrow B$. Users normally need not change the $defn$
setup.
Definitions may be presented with explicit arguments on the LHS, as
well as additional conditions, e.g.\ $f\;x\;y = t$ instead of $f
\equiv \lambda x\;y. t$ and $y \not= 0 \Imp g\;x\;y = u$ instead of
an unguarded $g \equiv \lambda x\;y. u$.
\item $\isarkeyword{abbreviation}~c~\isarkeyword{where}~eq$ introduces
a syntactic constant which is associated with a certain term
according to the meta-level equality $eq$.
Abbreviations participate in the usual type-inference process, but
are expanded before the logic ever sees them. Pretty printing of
terms involves higher-order rewriting with rules stemming from
reverted abbreviations. This needs some care to avoid overlapping
or looping syntactic replacements!
The optional $mode$ specification restricts output to a particular
print mode; using ``$input$'' here achieves the effect of one-way
abbreviations. The mode may also include an ``$output$'' qualifier
that affects the concrete syntax declared for abbreviations, cf.\
$\isarkeyword{syntax}$ in \S\ref{sec:syn-trans}.
\item $\isarkeyword{print_abbrevs}$ prints all constant abbreviations
of the current context.
\item $\isarkeyword{notation}~c~mx$ associates mixfix syntax with an
existing constant or fixed variable. This is a robust interface to
the underlying $\isarkeyword{syntax}$ primitive
(\S\ref{sec:syn-trans}). Type declaration and internal syntactic
representation of the given entity is retrieved from the context.
\item $\isarkeyword{no_notation}$ is similar to
$\isarkeyword{notation}$, but removes the specified syntax
annotation from the present context.
\end{descr}
All of these specifications support local theory targets (cf.\
\S\ref{sec:target}).
\subsection{Generic declarations}
Arbitrary operations on the background context may be wrapped-up as
generic declaration elements. Since the underlying concept of local
theories may be subject to later re-interpretation, there is an
additional dependency on a morphism that tells the difference of the
original declaration context wrt.\ the application context encountered
later on. A fact declaration is an important special case: it
consists of a theorem which is applied to the context by means of an
attribute.
\indexisarcmd{declaration}\indexisarcmd{declare}
\begin{matharray}{rcl}
\isarcmd{declaration} & : & \isarkeep{local{\dsh}theory} \\
\isarcmd{declare} & : & \isarkeep{local{\dsh}theory} \\
\end{matharray}
\begin{rail}
'declaration' target? text
;
'declare' target? (thmrefs + 'and')
;
\end{rail}
\begin{descr}
\item [$\isarkeyword{declaration}~d$] adds the declaration function
$d$ of ML type \verb,declaration, to the current local theory under
construction. In later application contexts, the function is
transformed according to the morphisms being involved in the
interpretation hierarchy.
\item [$\isarkeyword{declare}~thms$] declares theorems to the current
local theory context. 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}
\subsection{Local theory targets}\label{sec:target}
A local theory target is a context managed separately within the
enclosing theory. Contexts may introduce parameters (fixed variables)
and assumptions (hypotheses). Definitions and theorems depending on
the context may be added incrementally later on. Named contexts refer
to locales (cf.\ \S\ref{sec:locale}) or type classes (cf.\
\S\ref{sec:class}); the name ``$-$'' signifies the global theory
context.
\indexisarcmd{context}\indexisarcmd{end}
\begin{matharray}{rcll}
\isarcmd{context} & : & \isartrans{theory}{local{\dsh}theory} \\
\isarcmd{end} & : & \isartrans{local{\dsh}theory}{theory} \\
\end{matharray}
\indexouternonterm{target}
\begin{rail}
'context' name 'begin'
;
target: '(' 'in' name ')'
;
\end{rail}
\begin{descr}
\item $\isarkeyword{context}~c~\isarkeyword{begin}$ recommences an
existing locale or class context $c$. Note that locale and class
definitions allow to include the $\isarkeyword{begin}$ keyword as
well, in order to continue the local theory immediately after the
initial specification.
\item $\END$ concludes the current local theory and continues the
enclosing global theory. Note that a non-local $\END$ has a
different meaning: it concludes the theory itself
(\S\ref{sec:begin-thy}).
\item $(\IN~loc)$ given after any local theory command specifies an
immediate target, e.g.\
``$\isarkeyword{definition}~(\IN~loc)~\dots$'' or
``$\THEOREMNAME~(\IN~loc)~\dots$''. This works both in a local or
global theory context; the current target context will be suspended
for this command only. Note that $(\IN~-)$ will always produce a
global result independently of the current target context.
\end{descr}
The exact meaning of results produced within a local theory context
depends on the underlying target infrastructure (locale, type class
etc.). The general idea is as follows, considering a context named
$c$ with parameter $x$ and assumption $A[x]$.
Definitions are exported by introducing a global version with
additional arguments; a syntactic abbreviation links the long form
with the abstract version of the target context. For example, $a
\equiv t[x]$ becomes $c\dtt a \; ?x \equiv t[?x]$ at the theory level
(for arbitrary $?x$), together with a local abbreviation $c \equiv
c\dtt a\; x$ in the target context (for fixed $x$).
Theorems are exported by discharging the assumptions and generalizing
the parameters of the context. For example, $a: B[x]$ becomes $c\dtt
a: A[?x] \Imp B[?x]$ (for arbitrary $?x$).
\subsection{Locales}\label{sec:locale}
Locales are named local contexts, consisting of a list of declaration elements
that are modeled after the Isar proof context commands (cf.\
\S\ref{sec:proof-context}).
\subsubsection{Locale specifications}
\indexisarcmd{locale}\indexisarcmd{print-locale}\indexisarcmd{print-locales}
\begin{matharray}{rcl}
\isarcmd{locale} & : & \isartrans{theory}{local{\dsh}theory} \\
\isarcmd{print_locale}^* & : & \isarkeep{theory~|~proof} \\
\isarcmd{print_locales}^* & : & \isarkeep{theory~|~proof} \\
intro_locales & : & \isarmeth \\
unfold_locales & : & \isarmeth \\
\end{matharray}
\indexouternonterm{contextexpr}\indexouternonterm{contextelem}
\indexisarelem{fixes}\indexisarelem{constrains}\indexisarelem{assumes}
\indexisarelem{defines}\indexisarelem{notes}\indexisarelem{includes}
\begin{rail}
'locale' ('(open)')? name ('=' localeexpr)? 'begin'?
;
'print\_locale' '!'? localeexpr
;
localeexpr: ((contextexpr '+' (contextelem+)) | contextexpr | (contextelem+))
;
contextexpr: nameref | '(' contextexpr ')' |
(contextexpr (name mixfix? +)) | (contextexpr + '+')
;
contextelem: fixes | constrains | assumes | defines | notes | includes
;
fixes: 'fixes' ((name ('::' type)? structmixfix? | vars) + 'and')
;
constrains: 'constrains' (name '::' type + 'and')
;
assumes: 'assumes' (thmdecl? props + 'and')
;
defines: 'defines' (thmdecl? prop proppat? + 'and')
;
notes: 'notes' (thmdef? thmrefs + 'and')
;
includes: 'includes' contextexpr
;
\end{rail}
\begin{descr}
\item [$\LOCALE~loc~=~import~+~body$] defines a new locale $loc$ as a context
consisting of a certain view of existing locales ($import$) plus some
additional elements ($body$). Both $import$ and $body$ are optional; the
degenerate form $\LOCALE~loc$ defines an empty locale, which may still be
useful to collect declarations of facts later on. Type-inference on locale
expressions automatically takes care of the most general typing that the
combined context elements may acquire.
The $import$ consists of a structured context expression, consisting of
references to existing locales, renamed contexts, or merged contexts.
Renaming uses positional notation: $c~\vec x$ means that (a prefix of) the
fixed parameters of context $c$ are named according to $\vec x$; a
``\texttt{_}'' (underscore) \indexisarthm{_@\texttt{_}} means to skip that
position. Renaming by default deletes existing syntax. Optionally,
new syntax may by specified with a mixfix annotation. Note that the
special syntax declared with ``$(structure)$'' (see below) is
neither deleted nor can it be changed.
Merging proceeds from left-to-right, suppressing any duplicates stemming
from different paths through the import hierarchy.
The $body$ consists of basic context elements, further context expressions
may be included as well.
\begin{descr}
\item [$\FIXES{~x::\tau~(mx)}$] declares a local parameter of type $\tau$
and mixfix annotation $mx$ (both are optional). The special syntax
declaration ``$(structure)$'' means that $x$ may be referenced
implicitly in this context.
\item [$\CONSTRAINS{~x::\tau}$] introduces a type constraint $\tau$
on the local parameter $x$.
\item [$\ASSUMES{a}{\vec\phi}$] introduces local premises, similar to
$\ASSUMENAME$ within a proof (cf.\ \S\ref{sec:proof-context}).
\item [$\DEFINES{a}{x \equiv t}$] defines a previously declared parameter.
This is close to $\DEFNAME$ within a proof (cf.\
\S\ref{sec:proof-context}), but $\DEFINESNAME$ takes an equational
proposition instead of variable-term pair. The left-hand side of the
equation may have additional arguments, e.g.\ ``$\DEFINES{}{f~\vec x
\equiv t}$''.
\item [$\NOTES{a}{\vec b}$] reconsiders facts within a local context. Most
notably, this may include arbitrary declarations in any attribute
specifications included here, e.g.\ a local $simp$ rule.
\item [$\INCLUDES{c}$] copies the specified context in a statically scoped
manner. Only available in the long goal format of \S\ref{sec:goals}.
In contrast, the initial $import$ specification of a locale expression
maintains a dynamic relation to the locales being referenced (benefiting
from any later fact declarations in the obvious manner).
\end{descr}
Note that ``$\IS{p}$'' patterns given in the syntax of $\ASSUMESNAME$ and
$\DEFINESNAME$ above are illegal in locale definitions. In the long goal
format of \S\ref{sec:goals}, term bindings may be included as expected,
though.
\medskip By default, locale specifications are ``closed up'' by turning the
given text into a predicate definition $loc_axioms$ and deriving the
original assumptions as local lemmas (modulo local definitions). The
predicate statement covers only the newly specified assumptions, omitting
the content of included locale expressions. The full cumulative view is
only provided on export, involving another predicate $loc$ that refers to
the complete specification text.
In any case, the predicate arguments are those locale parameters that
actually occur in the respective piece of text. Also note that these
predicates operate at the meta-level in theory, but the locale packages
attempts to internalize statements according to the object-logic setup
(e.g.\ replacing $\Forall$ by $\forall$, and $\Imp$ by $\imp$ in HOL; see
also \S\ref{sec:object-logic}). Separate introduction rules
$loc_axioms.intro$ and $loc.intro$ are declared as well.
The $(open)$ option of a locale specification prevents both the current
$loc_axioms$ and cumulative $loc$ predicate constructions. Predicates are
also omitted for empty specification texts.
\item [$\isarkeyword{print_locale}~import~+~body$] prints the specified locale
expression in a flattened form. The notable special case
$\isarkeyword{print_locale}~loc$ just prints the contents of the named
locale, but keep in mind that type-inference will normalize type variables
according to the usual alphabetical order. The command omits
$\isarkeyword{notes}$ elements by default. Use
$\isarkeyword{print_locale}!$ to get them included.
\item [$\isarkeyword{print_locales}$] prints the names of all locales of the
current theory.
\item [$intro_locales$ and $unfold_locales$] repeatedly expand
all introduction rules of locale predicates of the theory. While
$intro_locales$ only applies the $loc.intro$ introduction rules and
therefore does not decend to assumptions, $unfold_locales$ is more
aggressive and applies $loc_axioms.intro$ as well. Both methods are
aware of locale specifications entailed by the context, both from
target and $\isarkeyword{includes}$ statements, and from
interpretations (see below). New goals that are entailed by the
current context are discharged automatically.
\end{descr}
\subsubsection{Interpretation of locales}
Locale expressions (more precisely, \emph{context expressions}) may be
instantiated, and the instantiated facts added to the current context.
This requires a proof of the instantiated specification and is called
\emph{locale interpretation}. Interpretation is possible in theories
and locales (command $\isarcmd{interpretation}$) and also in proof
contexts ($\isarcmd{interpret}$).
\indexisarcmd{interpretation}\indexisarcmd{interpret}
\indexisarcmd{print-interps}
\begin{matharray}{rcl}
\isarcmd{interpretation} & : & \isartrans{theory}{proof(prove)} \\
\isarcmd{interpret} & : & \isartrans{proof(state) ~|~ proof(chain)}{proof(prove)} \\
\isarcmd{print_interps}^* & : & \isarkeep{theory~|~proof} \\
\end{matharray}
\indexouternonterm{interp}
\railalias{printinterps}{print\_interps}
\railterm{printinterps}
\begin{rail}
'interpretation' (interp | name ('<' | subseteq) contextexpr)
;
'interpret' interp
;
printinterps '!'? name
;
instantiation: ('[' (inst+) ']')?
;
interp: thmdecl? \\ (contextexpr instantiation |
name instantiation 'where' (thmdecl? prop + 'and'))
;
\end{rail}
\begin{descr}
\item [$\isarcmd{interpretation}~expr~insts~\isarkeyword{where}~eqns$]
The first form of $\isarcmd{interpretation}$ interprets $expr$ in
the theory. The instantiation is given as a list of terms $insts$
and is positional. All parameters must receive an instantiation
term --- with the exception of defined parameters. These are, if
omitted, derived from the defining equation and other
instantiations. Use ``\_'' to omit an instantiation term. Free
variables are automatically generalized.
The command generates proof obligations for the instantiated
specifications (assumes and defines elements). Once these are
discharged by the user, instantiated facts are added to the theory in
a post-processing phase.
Additional equations, which are unfolded in facts during
post-processing, may be given after the keyword
$\isarkeyword{where}$. This is useful for interpreting concepts
introduced through definition specification elements. The equations
must be proved. Note that if equations are present, the context
expression is restricted to a locale name.
The command is aware of interpretations already active in the
theory. No proof obligations are generated for those, neither is
post-processing applied to their facts. This avoids duplication of
interpreted facts, in particular. Note that, in the case of a
locale with import, parts of the interpretation may already be
active. The command will only generate proof obligations and process
facts for new parts.
The context expression may be preceded by a name and/or attributes.
These take effect in the post-processing of facts. The name is used
to prefix fact names, for example to avoid accidental hiding of
other facts. Attributes are applied after attributes of the
interpreted facts.
Adding facts to locales has the
effect of adding interpreted facts to the theory for all active
interpretations also. That is, interpretations dynamically
participate in any facts added to locales.
\item [$\isarcmd{interpretation}~name~\subseteq~expr$]
This form of the command interprets $expr$ in the locale $name$. It
requires a proof that the specification of $name$ implies the
specification of $expr$. As in the localized version of the theorem
command, the proof is in the context of $name$. After the proof
obligation has been dischared, the facts of $expr$
become part of locale $name$ as \emph{derived} context elements and
are available when the context $name$ is subsequently entered.
Note that, like import, this is dynamic: facts added to a locale
part of $expr$ after interpretation become also available in
$name$. Like facts
of renamed context elements, facts obtained by interpretation may be
accessed by prefixing with the parameter renaming (where the parameters
are separated by `\_').
Unlike interpretation in theories, instantiation is confined to the
renaming of parameters, which may be specified as part of the context
expression $expr$. Using defined parameters in $name$ one may
achieve an effect similar to instantiation, though.
Only specification fragments of $expr$ that are not already part of
$name$ (be it imported, derived or a derived fragment of the import)
are considered by interpretation. This enables circular
interpretations.
If interpretations of $name$ exist in the current theory, the
command adds interpretations for $expr$ as well, with the same
prefix and attributes, although only for fragments of $expr$ that
are not interpreted in the theory already.
\item [$\isarcmd{interpret}~expr~insts~\isarkeyword{where}~eqns$]
interprets $expr$ in the proof context and is otherwise similar to
interpretation in theories. Free variables in instantiations are not
generalized, however.
\item [$\isarcmd{print_interps}~loc$]
prints the interpretations of a particular locale $loc$ that are
active in the current context, either theory or proof context. The
exclamation point argument triggers printing of
\emph{witness} theorems justifying interpretations. These are
normally omitted from the output.
\end{descr}
\begin{warn}
Since attributes are applied to interpreted theorems, interpretation
may modify the context of common proof tools, e.g.\ the Simplifier
or Classical Reasoner. Since the behavior of such automated
reasoning tools is \emph{not} stable under interpretation morphisms,
manual declarations might have to be issued.
\end{warn}
\begin{warn}
An interpretation in a theory may subsume previous interpretations.
This happens if the same specification fragment is interpreted twice
and the instantiation of the second interpretation is more general
than the interpretation of the first. A warning is issued, since it
is likely that these could have been generalized in the first place.
The locale package does not attempt to remove subsumed
interpretations.
\end{warn}
\subsection{Classes}\label{sec:class}
A class is a peculiarity of a locale with \emph{exactly one} type variable.
Beyond the underlying locale, a corresponding type class is established which
is interpreted logically as axiomatic type class \cite{Wenzel:1997:TPHOL}
whose logical content are the assumptions of the locale. Thus, classes provide
the full generality of locales combined with the commodity of type classes
(notably type-inference). See \cite{isabelle-classes} for a short tutorial.
\indexisarcmd{instance}\indexisarcmd{class}\indexisarcmd{print-classes}
\begin{matharray}{rcl}
\isarcmd{class} & : & \isartrans{theory}{local{\dsh}theory} \\
\isarcmd{instantiation} & : & \isartrans{theory}{local{\dsh}theory} \\
\isarcmd{instance} & : & \isartrans{local{\dsh}theory}{local{\dsh}theory} \\
\isarcmd{subclass} & : & \isartrans{local{\dsh}theory}{local{\dsh}theory} \\
\isarcmd{print_classes}^* & : & \isarkeep{theory~|~proof} \\
intro_classes & : & \isarmeth
\end{matharray}
\begin{rail}
'class' name '=' ((superclassexpr '+' (contextelem+)) | superclassexpr | (contextelem+)) \\
'begin'?
;
'instantiation' (nameref + 'and') '::' arity 'begin'
;
'instance'
;
'subclass' target? nameref
;
'print\_classes'
;
superclassexpr: nameref | (nameref '+' superclassexpr)
;
\end{rail}
\begin{descr}
\item [$\CLASS~c = superclasses~+~body$] defines a new class $c$,
inheriting from $superclasses$. This introduces a locale $c$
inheriting from all the locales $superclasses$. Correspondingly,
a type class $c$, inheriting from type classes $superclasses$.
$\FIXESNAME$ in $body$ are lifted to the global theory level
(\emph{class operations} $\vec f$ of class $c$),
mapping the local type parameter $\alpha$ to a schematic
type variable $?\alpha::c$
$\ASSUMESNAME$ in $body$ are also lifted, mapping each local parameter
$f::\tau [\alpha]$ to its corresponding global constant
$f::\tau [?\alpha::c]$.
\item [$\INSTANTIATION~\vec t~::~(\vec s)~s~\isarkeyword{begin}$]
opens a theory target (cf.\S\ref{sec:target}) which allows to specify
class operations $\vec f$ corresponding to sort $s$ at particular
type instances $\vec{\alpha::s}~t$ for each $t$ in $\vec t$.
An $\INSTANCE$ command in the target body sets up a goal stating
the type arities given after the $\INSTANTIATION$ keyword.
The possibility to give a list of type constructors with same arity
nicely corresponds to mutual recursive type definitions in Isabelle/HOL.
The target is concluded by an $\isarkeyword{end}$ keyword.
\item [$\INSTANCE$] in an instantiation target body sets up a goal stating
the type arities claimed at the opening $\INSTANTIATION$ keyword.
The proof would usually proceed by $intro_classes$, and then establish the
characteristic theorems of the type classes involved.
After finishing the proof, the background theory will be
augmented by the proven type arities.
\item [$\SUBCLASS~c$] in a class context for class $d$
sets up a goal stating that class $c$ is logically
contained in class $d$. After finishing the proof, class $d$ is proven
to be subclass $c$ and the locale $c$ is interpreted into $d$ simultaneously.
\item [$\isarkeyword{print_classes}$] prints all classes
in the current theory.
\item [$intro_classes$] repeatedly expands all class introduction rules of
this theory. Note that this method usually needs not be named explicitly,
as it is already included in the default proof step (of $\PROOFNAME$ etc.).
In particular, instantiation of trivial (syntactic) classes may be performed
by a single ``$\DDOT$'' proof step.
\end{descr}
\subsubsection{Class target}
A named context may refer to a locale (cf.~\S\ref{sec:target}). If this
locale is also a class $c$, beside the common locale target behaviour
the following occurs:
\begin{itemize}
\item Local constant declarations $g [\alpha]$ referring to the local type
parameter $\alpha$ and local parameters $\vec f [\alpha]$ are accompagnied
by theory-level constants $g [?\alpha::c]$ referring to theory-level
class operations $\vec f [?\alpha::c]$
\item Local theorem bindings are lifted as are assumptions.
\end{itemize}
\subsection{Axiomatic type classes}\label{sec:axclass}
\indexisarcmd{axclass}\indexisarmeth{intro-classes}
\begin{matharray}{rcl}
\isarcmd{axclass} & : & \isartrans{theory}{theory} \\
\isarcmd{instance} & : & \isartrans{theory}{proof(prove)} \\
\end{matharray}
Axiomatic type classes are Isabelle/Pure's primitive \emph{definitional} interface
to type classes. For practical applications, you should consider using classes
(cf.~\S\ref{sec:classes}) which provide a convenient user interface.
\begin{rail}
'axclass' classdecl (axmdecl prop +)
;
'instance' (nameref ('<' | subseteq) nameref | nameref '::' arity)
;
\end{rail}
\begin{descr}
\item [$\AXCLASS~c \subseteq \vec c~~axms$] defines an axiomatic type class as
the intersection of existing classes, with additional axioms holding. Class
axioms may not contain more than one type variable. The class axioms (with
implicit sort constraints added) are bound to the given names. Furthermore
a class introduction rule is generated (being bound as
$c_class{\dtt}intro$); this rule is employed by method $intro_classes$ to
support instantiation proofs of this class.
The ``axioms'' are stored as theorems according to the given name
specifications, adding the class name $c$ as name space prefix; the same
facts are also stored collectively as $c_class{\dtt}axioms$.
\item [$\INSTANCE~c@1 \subseteq c@2$ and $\INSTANCE~t :: (\vec s)s$] setup a
goal stating a class relation or type arity. The proof would usually
proceed by $intro_classes$, and then establish the characteristic theorems
of the type classes involved. After finishing the proof, the theory will be
augmented by a type signature declaration corresponding to the resulting
theorem.
\end{descr}
\subsection{Configuration options}
Isabelle/Pure maintains a record of named configuration options within the
theory or proof context, with values of type $bool$, $int$, or $string$.
Tools may declare options in ML, and then refer to these values (relative to
the context). Thus global reference variables are easily avoided. The user
may change the value of a configuration option by means of an associated
attribute of the same name. This form of context declaration works
particularly well with commands such as $\isarkeyword{declare}$ or
$\isarkeyword{using}$.
For historical reasons, some tools cannot take the full proof context
into account and merely refer to the background theory. This is
accommodated by configuration options being declared as ``global'',
which may not be changed within a local context.
\indexisarcmd{print-configs}
\begin{matharray}{rcll}
\isarcmd{print_configs} & : & \isarkeep{theory~|~proof} \\
\end{matharray}
\begin{rail}
name ('=' ('true' | 'false' | int | name))?
\end{rail}
\begin{descr}
\item [$\isarkeyword{print_configs}$] prints the available configuration
options, with names, types, and current values.
\item [$name = value$] as an attribute expression modifies the named option,
with the syntax of the value depending on the option's type. For $bool$ the
default value is $true$. Any attempt to change a global option in a local
context is ignored.
\end{descr}
\section{Derived proof schemes}
\subsection{Generalized elimination}\label{sec:obtain}
\indexisarcmd{obtain}\indexisarcmd{guess}
\begin{matharray}{rcl}
\isarcmd{obtain} & : & \isartrans{proof(state)}{proof(prove)} \\
\isarcmd{guess}^* & : & \isartrans{proof(state)}{proof(prove)} \\
\end{matharray}
Generalized elimination means that additional elements with certain properties
may be introduced in the current context, by virtue of a locally proven
``soundness statement''. Technically speaking, the $\OBTAINNAME$ language
element is like a declaration of $\FIXNAME$ and $\ASSUMENAME$ (see also see
\S\ref{sec:proof-context}), together with a soundness proof of its additional
claim. According to the nature of existential reasoning, assumptions get
eliminated from any result exported from the context later, provided that the
corresponding parameters do \emph{not} occur in the conclusion.
\begin{rail}
'obtain' parname? (vars + 'and') 'where' (props + 'and')
;
'guess' (vars + 'and')
;
\end{rail}
$\OBTAINNAME$ is defined as a derived Isar command as follows, where $\vec b$
shall refer to (optional) facts indicated for forward chaining.
\begin{matharray}{l}
\langle facts~\vec b\rangle \\
\OBTAIN{\vec x}{a}{\vec \phi}~~\langle proof\rangle \equiv {} \\[1ex]
\quad \HAVE{}{\All{thesis} (\All{\vec x} \vec\phi \Imp thesis) \Imp thesis} \\
\quad \PROOF{succeed} \\
\qquad \FIX{thesis} \\
\qquad \ASSUME{that~[intro?]}{\All{\vec x} \vec\phi \Imp thesis} \\
\qquad \THUS{}{thesis} \\
\quad\qquad \APPLY{-} \\
\quad\qquad \USING{\vec b}~~\langle proof\rangle \\
\quad \QED{} \\
\quad \FIX{\vec x}~\ASSUMENAME^\ast~a\colon~\vec\phi \\
\end{matharray}
Typically, the soundness proof is relatively straight-forward, often just by
canonical automated tools such as ``$\BY{simp}$'' or ``$\BY{blast}$''.
Accordingly, the ``$that$'' reduction above is declared as simplification and
introduction rule.
In a sense, $\OBTAINNAME$ represents at the level of Isar proofs what would be
meta-logical existential quantifiers and conjunctions. This concept has a
broad range of useful applications, ranging from plain elimination (or
introduction) of object-level existential and conjunctions, to elimination
over results of symbolic evaluation of recursive definitions, for example.
Also note that $\OBTAINNAME$ without parameters acts much like $\HAVENAME$,
where the result is treated as a genuine assumption.
An alternative name to be used instead of ``$that$'' above may be
given in parentheses.
\medskip
The improper variant $\isarkeyword{guess}$ is similar to $\OBTAINNAME$, but
derives the obtained statement from the course of reasoning! The proof starts
with a fixed goal $thesis$. The subsequent proof may refine this to anything
of the form like $\All{\vec x} \vec\phi \Imp thesis$, but must not introduce
new subgoals. The final goal state is then used as reduction rule for the
obtain scheme described above. Obtained parameters $\vec x$ are marked as
internal by default, which prevents the proof context from being polluted by
ad-hoc variables. The variable names and type constraints given as arguments
for $\isarkeyword{guess}$ specify a prefix of obtained parameters explicitly
in the text.
It is important to note that the facts introduced by $\OBTAINNAME$ and
$\isarkeyword{guess}$ may not be polymorphic: any type-variables occurring
here are fixed in the present context!
\subsection{Calculational reasoning}\label{sec:calculation}
\indexisarcmd{also}\indexisarcmd{finally}
\indexisarcmd{moreover}\indexisarcmd{ultimately}
\indexisarcmd{print-trans-rules}
\indexisaratt{trans}\indexisaratt{sym}\indexisaratt{symmetric}
\begin{matharray}{rcl}
\isarcmd{also} & : & \isartrans{proof(state)}{proof(state)} \\
\isarcmd{finally} & : & \isartrans{proof(state)}{proof(chain)} \\
\isarcmd{moreover} & : & \isartrans{proof(state)}{proof(state)} \\
\isarcmd{ultimately} & : & \isartrans{proof(state)}{proof(chain)} \\
\isarcmd{print_trans_rules}^* & : & \isarkeep{theory~|~proof} \\
trans & : & \isaratt \\
sym & : & \isaratt \\
symmetric & : & \isaratt \\
\end{matharray}
Calculational proof is forward reasoning with implicit application of
transitivity rules (such those of $=$, $\leq$, $<$). Isabelle/Isar maintains
an auxiliary register $calculation$\indexisarthm{calculation} for accumulating
results obtained by transitivity composed with the current result. Command
$\ALSO$ updates $calculation$ involving $this$, while $\FINALLY$ exhibits the
final $calculation$ by forward chaining towards the next goal statement. Both
commands require valid current facts, i.e.\ may occur only after commands that
produce theorems such as $\ASSUMENAME$, $\NOTENAME$, or some finished proof of
$\HAVENAME$, $\SHOWNAME$ etc. The $\MOREOVER$ and $\ULTIMATELY$ commands are
similar to $\ALSO$ and $\FINALLY$, but only collect further results in
$calculation$ without applying any rules yet.
Also note that the implicit term abbreviation ``$\dots$'' has its canonical
application with calculational proofs. It refers to the argument of the
preceding statement. (The argument of a curried infix expression happens to be
its right-hand side.)
Isabelle/Isar calculations are implicitly subject to block structure in the
sense that new threads of calculational reasoning are commenced for any new
block (as opened by a local goal, for example). This means that, apart from
being able to nest calculations, there is no separate \emph{begin-calculation}
command required.
\medskip
The Isar calculation proof commands may be defined as follows:\footnote{We
suppress internal bookkeeping such as proper handling of block-structure.}
\begin{matharray}{rcl}
\ALSO@0 & \equiv & \NOTE{calculation}{this} \\
\ALSO@{n+1} & \equiv & \NOTE{calculation}{trans~[OF~calculation~this]} \\[0.5ex]
\FINALLY & \equiv & \ALSO~\FROM{calculation} \\
\MOREOVER & \equiv & \NOTE{calculation}{calculation~this} \\
\ULTIMATELY & \equiv & \MOREOVER~\FROM{calculation} \\
\end{matharray}
\begin{rail}
('also' | 'finally') ('(' thmrefs ')')?
;
'trans' (() | 'add' | 'del')
;
\end{rail}
\begin{descr}
\item [$\ALSO~(\vec a)$] maintains the auxiliary $calculation$ register as
follows. The first occurrence of $\ALSO$ in some calculational thread
initializes $calculation$ by $this$. Any subsequent $\ALSO$ on the same
level of block-structure updates $calculation$ by some transitivity rule
applied to $calculation$ and $this$ (in that order). Transitivity rules are
picked from the current context, unless alternative rules are given as
explicit arguments.
\item [$\FINALLY~(\vec a)$] maintaining $calculation$ in the same way as
$\ALSO$, and concludes the current calculational thread. The final result
is exhibited as fact for forward chaining towards the next goal. Basically,
$\FINALLY$ just abbreviates $\ALSO~\FROM{calculation}$. Note that
``$\FINALLY~\SHOW{}{\Var{thesis}}~\DOT$'' and
``$\FINALLY~\HAVE{}{\phi}~\DOT$'' are typical idioms for concluding
calculational proofs.
\item [$\MOREOVER$ and $\ULTIMATELY$] are analogous to $\ALSO$ and $\FINALLY$,
but collect results only, without applying rules.
\item [$\isarkeyword{print_trans_rules}$] prints the list of transitivity
rules (for calculational commands $\ALSO$ and $\FINALLY$) and symmetry rules
(for the $symmetric$ operation and single step elimination patters) of the
current context.
\item [$trans$] declares theorems as transitivity rules.
\item [$sym$] declares symmetry rules.
\item [$symmetric$] resolves a theorem with some rule declared as $sym$ in the
current context. For example, ``$\ASSUME{[symmetric]}{x = y}$'' produces a
swapped fact derived from that assumption.
In structured proof texts it is often more appropriate to use an explicit
single-step elimination proof, such as ``$\ASSUME{}{x = y}~\HENCE{}{y =
x}~\DDOT$''. The very same rules known to $symmetric$ are declared as
$elim?$ as well.
\end{descr}
\section{Proof tools}
\subsection{Miscellaneous methods and attributes}\label{sec:misc-meth-att}
\indexisarmeth{unfold}\indexisarmeth{fold}\indexisarmeth{insert}
\indexisarmeth{erule}\indexisarmeth{drule}\indexisarmeth{frule}
\indexisarmeth{fail}\indexisarmeth{succeed}
\begin{matharray}{rcl}
unfold & : & \isarmeth \\
fold & : & \isarmeth \\
insert & : & \isarmeth \\[0.5ex]
erule^* & : & \isarmeth \\
drule^* & : & \isarmeth \\
frule^* & : & \isarmeth \\
succeed & : & \isarmeth \\
fail & : & \isarmeth \\
\end{matharray}
\begin{rail}
('fold' | 'unfold' | 'insert') thmrefs
;
('erule' | 'drule' | 'frule') ('('nat')')? thmrefs
;
\end{rail}
\begin{descr}
\item [$unfold~\vec a$ and $fold~\vec a$] expand (or fold back again)
the given definitions throughout all goals; any chained facts
provided are inserted into the goal and subject to rewriting as
well.
\item [$insert~\vec a$] inserts theorems as facts into all goals of the proof
state. Note that current facts indicated for forward chaining are ignored.
\item [$erule~\vec a$, $drule~\vec a$, and $frule~\vec a$] are similar to the
basic $rule$ method (see \S\ref{sec:pure-meth-att}), but apply rules by
elim-resolution, destruct-resolution, and forward-resolution, respectively
\cite{isabelle-ref}. The optional natural number argument (default $0$)
specifies additional assumption steps to be performed here.
Note that these methods are improper ones, mainly serving for
experimentation and tactic script emulation. Different modes of basic rule
application are usually expressed in Isar at the proof language level,
rather than via implicit proof state manipulations. For example, a proper
single-step elimination would be done using the plain $rule$ method, with
forward chaining of current facts.
\item [$succeed$] yields a single (unchanged) result; it is the identity of
the ``\texttt{,}'' method combinator (cf.\ \S\ref{sec:syn-meth}).
\item [$fail$] yields an empty result sequence; it is the identity of the
``\texttt{|}'' method combinator (cf.\ \S\ref{sec:syn-meth}).
\end{descr}
\indexisaratt{tagged}\indexisaratt{untagged}
\indexisaratt{THEN}\indexisaratt{COMP}
\indexisaratt{unfolded}\indexisaratt{folded}
\indexisaratt{standard}\indexisarattof{Pure}{elim-format}
\indexisaratt{no-vars}
\begin{matharray}{rcl}
tagged & : & \isaratt \\
untagged & : & \isaratt \\[0.5ex]
THEN & : & \isaratt \\
COMP & : & \isaratt \\[0.5ex]
unfolded & : & \isaratt \\
folded & : & \isaratt \\[0.5ex]
rotated & : & \isaratt \\
elim_format & : & \isaratt \\
standard^* & : & \isaratt \\
no_vars^* & : & \isaratt \\
\end{matharray}
\begin{rail}
'tagged' nameref
;
'untagged' name
;
('THEN' | 'COMP') ('[' nat ']')? thmref
;
('unfolded' | 'folded') thmrefs
;
'rotated' ( int )?
\end{rail}
\begin{descr}
\item [$tagged~name~arg$ and $untagged~name$] add and remove $tags$ of some
theorem. Tags may be any list of strings that serve as comment for some
tools (e.g.\ $\LEMMANAME$ causes the tag ``$lemma$'' to be added to the
result). The first string is considered the tag name, the second its
argument. Note that $untagged$ removes any tags of the same name.
\item [$THEN~a$ and $COMP~a$] compose rules by resolution. $THEN$ resolves
with the first premise of $a$ (an alternative position may be also
specified); the $COMP$ version skips the automatic lifting process that is
normally intended (cf.\ \texttt{RS} and \texttt{COMP} in
\cite[\S5]{isabelle-ref}).
\item [$unfolded~\vec a$ and $folded~\vec a$] expand and fold back
again the given definitions throughout a rule.
\item [$rotated~n$] rotate the premises of a theorem by $n$ (default 1).
\item [$elim_format$] turns a destruction rule into elimination rule format,
by resolving with the rule $\PROP A \Imp (\PROP A \Imp \PROP B) \Imp \PROP
B$.
Note that the Classical Reasoner (\S\ref{sec:classical}) provides its own
version of this operation.
\item [$standard$] puts a theorem into the standard form of object-rules at
the outermost theory level. Note that this operation violates the local
proof context (including active locales).
\item [$no_vars$] replaces schematic variables by free ones; this is mainly
for tuning output of pretty printed theorems.
\end{descr}
\subsection{Further tactic emulations}\label{sec:tactics}
The following improper proof methods emulate traditional tactics. These admit
direct access to the goal state, which is normally considered harmful! In
particular, this may involve both numbered goal addressing (default 1), and
dynamic instantiation within the scope of some subgoal.
\begin{warn}
Dynamic instantiations refer to universally quantified parameters of
a subgoal (the dynamic context) rather than fixed variables and term
abbreviations of a (static) Isar context.
\end{warn}
Tactic emulation methods, unlike their ML counterparts, admit
simultaneous instantiation from both dynamic and static contexts. If
names occur in both contexts goal parameters hide locally fixed
variables. Likewise, schematic variables refer to term abbreviations,
if present in the static context. Otherwise the schematic variable is
interpreted as a schematic variable and left to be solved by unification
with certain parts of the subgoal.
Note that the tactic emulation proof methods in Isabelle/Isar are consistently
named $foo_tac$. Note also that variable names occurring on left hand sides
of instantiations must be preceded by a question mark if they coincide with
a keyword or contain dots.
This is consistent with the attribute $where$ (see \S\ref{sec:pure-meth-att}).
\indexisarmeth{rule-tac}\indexisarmeth{erule-tac}
\indexisarmeth{drule-tac}\indexisarmeth{frule-tac}
\indexisarmeth{cut-tac}\indexisarmeth{thin-tac}
\indexisarmeth{subgoal-tac}\indexisarmeth{rename-tac}
\indexisarmeth{rotate-tac}\indexisarmeth{tactic}
\begin{matharray}{rcl}
rule_tac^* & : & \isarmeth \\
erule_tac^* & : & \isarmeth \\
drule_tac^* & : & \isarmeth \\
frule_tac^* & : & \isarmeth \\
cut_tac^* & : & \isarmeth \\
thin_tac^* & : & \isarmeth \\
subgoal_tac^* & : & \isarmeth \\
rename_tac^* & : & \isarmeth \\
rotate_tac^* & : & \isarmeth \\
tactic^* & : & \isarmeth \\
\end{matharray}
\railalias{ruletac}{rule\_tac}
\railterm{ruletac}
\railalias{eruletac}{erule\_tac}
\railterm{eruletac}
\railalias{druletac}{drule\_tac}
\railterm{druletac}
\railalias{fruletac}{frule\_tac}
\railterm{fruletac}
\railalias{cuttac}{cut\_tac}
\railterm{cuttac}
\railalias{thintac}{thin\_tac}
\railterm{thintac}
\railalias{subgoaltac}{subgoal\_tac}
\railterm{subgoaltac}
\railalias{renametac}{rename\_tac}
\railterm{renametac}
\railalias{rotatetac}{rotate\_tac}
\railterm{rotatetac}
\begin{rail}
( ruletac | eruletac | druletac | fruletac | cuttac | thintac ) goalspec?
( insts thmref | thmrefs )
;
subgoaltac goalspec? (prop +)
;
renametac goalspec? (name +)
;
rotatetac goalspec? int?
;
'tactic' text
;
insts: ((name '=' term) + 'and') 'in'
;
\end{rail}
\begin{descr}
\item [$rule_tac$ etc.] do resolution of rules with explicit instantiation.
This works the same way as the ML tactics \texttt{res_inst_tac} etc. (see
\cite[\S3]{isabelle-ref}).
Multiple rules may be only given if there is no instantiation; then
$rule_tac$ is the same as \texttt{resolve_tac} in ML (see
\cite[\S3]{isabelle-ref}).
\item [$cut_tac$] inserts facts into the proof state as assumption of a
subgoal, see also \texttt{cut_facts_tac} in \cite[\S3]{isabelle-ref}. Note
that the scope of schematic variables is spread over the main goal
statement. Instantiations may be given as well, see also ML tactic
\texttt{cut_inst_tac} in \cite[\S3]{isabelle-ref}.
\item [$thin_tac~\phi$] deletes the specified assumption from a subgoal; note
that $\phi$ may contain schematic variables. See also \texttt{thin_tac} in
\cite[\S3]{isabelle-ref}.
\item [$subgoal_tac~\phi$] adds $\phi$ as an assumption to a subgoal. See
also \texttt{subgoal_tac} and \texttt{subgoals_tac} in
\cite[\S3]{isabelle-ref}.
\item [$rename_tac~\vec x$] renames parameters of a goal according to the list
$\vec x$, which refers to the \emph{suffix} of variables.
\item [$rotate_tac~n$] rotates the assumptions of a goal by $n$ positions:
from right to left if $n$ is positive, and from left to right if $n$ is
negative; the default value is $1$. See also \texttt{rotate_tac} in
\cite[\S3]{isabelle-ref}.
\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:
{\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.
\end{descr}
\subsection{The Simplifier}\label{sec:simplifier}
\subsubsection{Simplification methods}
\indexisarmeth{simp}\indexisarmeth{simp-all}
\begin{matharray}{rcl}
simp & : & \isarmeth \\
simp_all & : & \isarmeth \\
\end{matharray}
\indexouternonterm{simpmod}
\begin{rail}
('simp' | 'simp\_all') ('!' ?) opt? (simpmod *)
;
opt: '(' ('no\_asm' | 'no\_asm\_simp' | 'no\_asm\_use' | 'asm\_lr' | 'depth\_limit' ':' nat) ')'
;
simpmod: ('add' | 'del' | 'only' | 'cong' (() | 'add' | 'del') |
'split' (() | 'add' | 'del')) ':' thmrefs
;
\end{rail}
\begin{descr}
\item [$simp$] invokes Isabelle's simplifier, after declaring additional rules
according to the arguments given. Note that the \railtterm{only} modifier
first removes all other rewrite rules, congruences, and looper tactics
(including splits), and then behaves like \railtterm{add}.
\medskip The \railtterm{cong} modifiers add or delete Simplifier congruence
rules (see also \cite{isabelle-ref}), the default is to add.
\medskip The \railtterm{split} modifiers add or delete rules for the
Splitter (see also \cite{isabelle-ref}), the default is to add. This works
only if the Simplifier method has been properly setup to include the
Splitter (all major object logics such HOL, HOLCF, FOL, ZF do this already).
\item [$simp_all$] is similar to $simp$, but acts on all goals (backwards from
the last to the first one).
\end{descr}
By default the Simplifier methods take local assumptions fully into account,
using equational assumptions in the subsequent normalization process, or
simplifying assumptions themselves (cf.\ \texttt{asm_full_simp_tac} in
\cite[\S10]{isabelle-ref}). In structured proofs this is usually quite well
behaved in practice: just the local premises of the actual goal are involved,
additional facts may be inserted via explicit forward-chaining (using $\THEN$,
$\FROMNAME$ etc.). The full context of assumptions is only included if the
``$!$'' (bang) argument is given, which should be used with some care, though.
Additional Simplifier options may be specified to tune the behavior further
(mostly for unstructured scripts with many accidental local facts):
``$(no_asm)$'' means assumptions are ignored completely (cf.\
\texttt{simp_tac}), ``$(no_asm_simp)$'' means assumptions are used in the
simplification of the conclusion but are not themselves simplified (cf.\
\texttt{asm_simp_tac}), and ``$(no_asm_use)$'' means assumptions are
simplified but are not used in the simplification of each other or the
conclusion (cf.\ \texttt{full_simp_tac}). For compatibility reasons, there is
also an option ``$(asm_lr)$'', which means that an assumption is only used for
simplifying assumptions which are to the right of it (cf.\
\texttt{asm_lr_simp_tac}). Giving an option ``$(depth_limit: n)$'' limits the
number of recursive invocations of the simplifier during conditional
rewriting.
\medskip
The Splitter package is usually configured to work as part of the Simplifier.
The effect of repeatedly applying \texttt{split_tac} can be simulated by
``$(simp~only\colon~split\colon~\vec a)$''. There is also a separate $split$
method available for single-step case splitting.
\subsubsection{Declaring rules}
\indexisarcmd{print-simpset}
\indexisaratt{simp}\indexisaratt{split}\indexisaratt{cong}
\begin{matharray}{rcl}
\isarcmd{print_simpset}^* & : & \isarkeep{theory~|~proof} \\
simp & : & \isaratt \\
cong & : & \isaratt \\
split & : & \isaratt \\
\end{matharray}
\begin{rail}
('simp' | 'cong' | 'split') (() | 'add' | 'del')
;
\end{rail}
\begin{descr}
\item [$\isarcmd{print_simpset}$] prints the collection of rules declared to
the Simplifier, which is also known as ``simpset'' internally
\cite{isabelle-ref}. This is a diagnostic command; $undo$ does not apply.
\item [$simp$] declares simplification rules.
\item [$cong$] declares congruence rules.
\item [$split$] declares case split rules.
\end{descr}
\subsubsection{Simplification procedures}
\indexisarcmd{simproc-setup}
\indexisaratt{simproc}
\begin{matharray}{rcl}
\isarcmd{simproc_setup} & : & \isarkeep{local{\dsh}theory} \\
simproc & : & \isaratt \\
\end{matharray}
\begin{rail}
'simproc\_setup' name '(' (term + '|') ')' '=' text \\ ('identifier' (nameref+))?
;
'simproc' (('add' ':')? | 'del' ':') (name+)
;
\end{rail}
\begin{descr}
\item [$\isarcmd{simproc_setup}$] defines a named simplification
procedure that is invoked by the Simplifier whenever any of the
given term patterns match the current redex. The implementation,
which is provided as ML source text, needs to be of type
\verb,morphism -> simpset -> cterm -> thm option,, where the
\verb,cterm, represents the current redex $r$ and the result is
supposed to be some proven rewrite rule $r \equiv r'$ (or a
generalized version), or \verb,NONE, to indicate failure. The
\verb,simpset, argument holds the full context of the current
Simplifier invocation, including the actual Isar proof context. The
\verb,morphism, informs about the difference of the original
compilation context wrt.\ the one of the actual application later
on. The optional $\isarkeyword{identifier}$ specifies theorems that
represent the logical content of the abstract theory of this
simproc.
Morphisms and identifiers are only relevant for simprocs that are
defined within a local target context, e.g.\ in a locale.
\item [$simproc\;add\colon\;name$ and $simproc\;del\colon\;name$] add
or delete named simprocs to the current Simplifier context. The
default is to add a simproc. Note that $\isarcmd{simproc_setup}$
already adds the new simproc to the subsequent context.
\end{descr}
\subsubsection{Forward simplification}
\indexisaratt{simplified}
\begin{matharray}{rcl}
simplified & : & \isaratt \\
\end{matharray}
\begin{rail}
'simplified' opt? thmrefs?
;
opt: '(' (noasm | noasmsimp | noasmuse) ')'
;
\end{rail}
\begin{descr}
\item [$simplified~\vec a$] causes a theorem to be simplified, either by
exactly the specified rules $\vec a$, or the implicit Simplifier context if
no arguments are given. The result is fully simplified by default,
including assumptions and conclusion; the options $no_asm$ etc.\ tune the
Simplifier in the same way as the for the $simp$ method.
Note that forward simplification restricts the simplifier to its most basic
operation of term rewriting; solver and looper tactics \cite{isabelle-ref}
are \emph{not} involved here. The $simplified$ attribute should be only
rarely required under normal circumstances.
\end{descr}
\subsubsection{Low-level equational reasoning}
\indexisarmeth{subst}\indexisarmeth{hypsubst}\indexisarmeth{split}
\begin{matharray}{rcl}
subst^* & : & \isarmeth \\
hypsubst^* & : & \isarmeth \\
split^* & : & \isarmeth \\
\end{matharray}
\begin{rail}
'subst' ('(' 'asm' ')')? ('(' (nat+) ')')? thmref
;
'split' ('(' 'asm' ')')? thmrefs
;
\end{rail}
These methods provide low-level facilities for equational reasoning that are
intended for specialized applications only. Normally, single step
calculations would be performed in a structured text (see also
\S\ref{sec:calculation}), while the Simplifier methods provide the canonical
way for automated normalization (see \S\ref{sec:simplifier}).
\begin{descr}
\item [$subst~eq$] performs a single substitution step using rule $eq$, which
may be either a meta or object equality.
\item [$subst~(asm)~eq$] substitutes in an assumption.
\item [$subst~(i \dots j)~eq$] performs several substitutions in the
conclusion. The numbers $i$ to $j$ indicate the positions to substitute at.
Positions are ordered from the top of the term tree moving down from left to
right. For example, in $(a+b)+(c+d)$ there are three positions where
commutativity of $+$ is applicable: 1 refers to the whole term, 2 to $a+b$
and 3 to $c+d$. If the positions in the list $(i \dots j)$ are
non-overlapping (e.g. $(2~3)$ in $(a+b)+(c+d)$) you may assume all
substitutions are performed simultaneously. Otherwise the behaviour of
$subst$ is not specified.
\item [$subst~(asm)~(i \dots j)~eq$] performs the substitutions in the
assumptions. Positions $1 \dots i@1$ refer
to assumption 1, positions $i@1+1 \dots i@2$ to assumption 2, and so on.
\item [$hypsubst$] performs substitution using some assumption; this only
works for equations of the form $x = t$ where $x$ is a free or bound
variable.
\item [$split~\vec a$] performs single-step case splitting using rules $thms$.
By default, splitting is performed in the conclusion of a goal; the $asm$
option indicates to operate on assumptions instead.
Note that the $simp$ method already involves repeated application of split
rules as declared in the current context.
\end{descr}
\subsection{The Classical Reasoner}\label{sec:classical}
\subsubsection{Basic methods}
\indexisarmeth{rule}\indexisarmeth{default}\indexisarmeth{contradiction}
\indexisarmeth{intro}\indexisarmeth{elim}
\begin{matharray}{rcl}
rule & : & \isarmeth \\
contradiction & : & \isarmeth \\
intro & : & \isarmeth \\
elim & : & \isarmeth \\
\end{matharray}
\begin{rail}
('rule' | 'intro' | 'elim') thmrefs?
;
\end{rail}
\begin{descr}
\item [$rule$] as offered by the classical reasoner is a refinement over the
primitive one (see \S\ref{sec:pure-meth-att}). Both versions essentially
work the same, but the classical version observes the classical rule context
in addition to that of Isabelle/Pure.
Common object logics (HOL, ZF, etc.) declare a rich collection of classical
rules (even if these would qualify as intuitionistic ones), but only few
declarations to the rule context of Isabelle/Pure
(\S\ref{sec:pure-meth-att}).
\item [$contradiction$] solves some goal by contradiction, deriving any result
from both $\lnot A$ and $A$. Chained facts, which are guaranteed to
participate, may appear in either order.
\item [$intro$ and $elim$] repeatedly refine some goal by intro- or
elim-resolution, after having inserted any chained facts. Exactly the rules
given as arguments are taken into account; this allows fine-tuned
decomposition of a proof problem, in contrast to common automated tools.
\end{descr}
\subsubsection{Automated methods}
\indexisarmeth{blast}\indexisarmeth{fast}\indexisarmeth{slow}
\indexisarmeth{best}\indexisarmeth{safe}\indexisarmeth{clarify}
\begin{matharray}{rcl}
blast & : & \isarmeth \\
fast & : & \isarmeth \\
slow & : & \isarmeth \\
best & : & \isarmeth \\
safe & : & \isarmeth \\
clarify & : & \isarmeth \\
\end{matharray}
\indexouternonterm{clamod}
\begin{rail}
'blast' ('!' ?) nat? (clamod *)
;
('fast' | 'slow' | 'best' | 'safe' | 'clarify') ('!' ?) (clamod *)
;
clamod: (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' thmrefs
;
\end{rail}
\begin{descr}
\item [$blast$] refers to the classical tableau prover (see \texttt{blast_tac}
in \cite[\S11]{isabelle-ref}). The optional argument specifies a
user-supplied search bound (default 20).
\item [$fast$, $slow$, $best$, $safe$, and $clarify$] refer to the generic
classical reasoner. See \texttt{fast_tac}, \texttt{slow_tac},
\texttt{best_tac}, \texttt{safe_tac}, and \texttt{clarify_tac} in
\cite[\S11]{isabelle-ref} for more information.
\end{descr}
Any of the above methods support additional modifiers of the context of
classical rules. Their semantics is analogous to the attributes given before.
Facts provided by forward chaining are inserted into the goal before
commencing proof search. The ``!''~argument causes the full context of
assumptions to be included as well.
\subsubsection{Combined automated methods}\label{sec:clasimp}
\indexisarmeth{auto}\indexisarmeth{force}\indexisarmeth{clarsimp}
\indexisarmeth{fastsimp}\indexisarmeth{slowsimp}\indexisarmeth{bestsimp}
\begin{matharray}{rcl}
auto & : & \isarmeth \\
force & : & \isarmeth \\
clarsimp & : & \isarmeth \\
fastsimp & : & \isarmeth \\
slowsimp & : & \isarmeth \\
bestsimp & : & \isarmeth \\
\end{matharray}
\indexouternonterm{clasimpmod}
\begin{rail}
'auto' '!'? (nat nat)? (clasimpmod *)
;
('force' | 'clarsimp' | 'fastsimp' | 'slowsimp' | 'bestsimp') '!'? (clasimpmod *)
;
clasimpmod: ('simp' (() | 'add' | 'del' | 'only') |
('cong' | 'split') (() | 'add' | 'del') |
'iff' (((() | 'add') '?'?) | 'del') |
(('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' thmrefs
\end{rail}
\begin{descr}
\item [$auto$, $force$, $clarsimp$, $fastsimp$, $slowsimp$, and $bestsimp$]
provide access to Isabelle's combined simplification and classical reasoning
tactics. These correspond to \texttt{auto_tac}, \texttt{force_tac},
\texttt{clarsimp_tac}, and Classical Reasoner tactics with the Simplifier
added as wrapper, see \cite[\S11]{isabelle-ref} for more information. The
modifier arguments correspond to those given in \S\ref{sec:simplifier} and
\S\ref{sec:classical}. Just note that the ones related to the Simplifier
are prefixed by \railtterm{simp} here.
Facts provided by forward chaining are inserted into the goal before doing
the search. The ``!''~argument causes the full context of assumptions to be
included as well.
\end{descr}
\subsubsection{Declaring rules}
\indexisarcmd{print-claset}
\indexisaratt{intro}\indexisaratt{elim}\indexisaratt{dest}
\indexisaratt{iff}\indexisaratt{rule}
\begin{matharray}{rcl}
\isarcmd{print_claset}^* & : & \isarkeep{theory~|~proof} \\
intro & : & \isaratt \\
elim & : & \isaratt \\
dest & : & \isaratt \\
rule & : & \isaratt \\
iff & : & \isaratt \\
\end{matharray}
\begin{rail}
('intro' | 'elim' | 'dest') ('!' | () | '?') nat?
;
'rule' 'del'
;
'iff' (((() | 'add') '?'?) | 'del')
;
\end{rail}
\begin{descr}
\item [$\isarcmd{print_claset}$] prints the collection of rules declared to
the Classical Reasoner, which is also known as ``claset'' internally
\cite{isabelle-ref}. This is a diagnostic command; $undo$ does not apply.
\item [$intro$, $elim$, and $dest$] declare introduction, elimination, and
destruction rules, respectively. By default, rules are considered as
\emph{unsafe} (i.e.\ not applied blindly without backtracking), while a
single ``!'' classifies as \emph{safe}. Rule declarations marked by ``?''
coincide with those of Isabelle/Pure, cf.\ \S\ref{sec:pure-meth-att} (i.e.\
are only applied in single steps of the $rule$ method). The optional
natural number specifies an explicit weight argument, which is ignored by
automated tools, but determines the search order of single rule steps.
\item [$rule~del$] deletes introduction, elimination, or destruction rules from
the context.
\item [$iff$] declares logical equivalences to the Simplifier and the
Classical reasoner at the same time. Non-conditional rules result in a
``safe'' introduction and elimination pair; conditional ones are considered
``unsafe''. Rules with negative conclusion are automatically inverted
(using $\lnot$ elimination internally).
The ``?'' version of $iff$ declares rules to the Isabelle/Pure context only,
and omits the Simplifier declaration.
\end{descr}
\subsubsection{Classical operations}
\indexisaratt{swapped}
\begin{matharray}{rcl}
swapped & : & \isaratt \\
\end{matharray}
\begin{descr}
\item [$swapped$] turns an introduction rule into an elimination, by resolving
with the classical swap principle $(\lnot B \Imp A) \Imp (\lnot A \Imp B)$.
\end{descr}
\subsection{Proof by cases and induction}\label{sec:cases-induct}
\subsubsection{Rule contexts}
\indexisarcmd{case}\indexisarcmd{print-cases}
\indexisaratt{case-names}\indexisaratt{case-conclusion}
\indexisaratt{params}\indexisaratt{consumes}
\begin{matharray}{rcl}
\isarcmd{case} & : & \isartrans{proof(state)}{proof(state)} \\
\isarcmd{print_cases}^* & : & \isarkeep{proof} \\
case_names & : & \isaratt \\
case_conclusion & : & \isaratt \\
params & : & \isaratt \\
consumes & : & \isaratt \\
\end{matharray}
The puristic way to build up Isar proof contexts is by explicit language
elements like $\FIXNAME$, $\ASSUMENAME$, $\LET$ (see
\S\ref{sec:proof-context}). This is adequate for plain natural deduction, but
easily becomes unwieldy in concrete verification tasks, which typically
involve big induction rules with several cases.
The $\CASENAME$ command provides a shorthand to refer to a local context
symbolically: certain proof methods provide an environment of named ``cases''
of the form $c\colon \vec x, \vec \phi$; the effect of ``$\CASE{c}$'' is then
equivalent to ``$\FIX{\vec x}~\ASSUME{c}{\vec\phi}$''. Term bindings may be
covered as well, notably $\Var{case}$ for the main conclusion.
By default, the ``terminology'' $\vec x$ of a case value is marked as hidden,
i.e.\ there is no way to refer to such parameters in the subsequent proof
text. After all, original rule parameters stem from somewhere outside of the
current proof text. By using the explicit form ``$\CASE{(c~\vec y)}$''
instead, the proof author is able to chose local names that fit nicely into
the current context.
\medskip
It is important to note that proper use of $\CASENAME$ does not provide means
to peek at the current goal state, which is not directly observable in Isar!
Nonetheless, goal refinement commands do provide named cases $goal@i$ for each
subgoal $i = 1, \dots, n$ of the resulting goal state. Using this feature
requires great care, because some bits of the internal tactical machinery
intrude the proof text. In particular, parameter names stemming from the
left-over of automated reasoning tools are usually quite unpredictable.
Under normal circumstances, the text of cases emerge from standard elimination
or induction rules, which in turn are derived from previous theory
specifications in a canonical way (say from $\isarkeyword{inductive}$
definitions).
\medskip Proper cases are only available if both the proof method and the
rules involved support this. By using appropriate attributes, case names,
conclusions, and parameters may be also declared by hand. Thus variant
versions of rules that have been derived manually become ready to use in
advanced case analysis later.
\begin{rail}
'case' (caseref | '(' caseref ((name | underscore) +) ')')
;
caseref: nameref attributes?
;
'case\_names' (name +)
;
'case\_conclusion' name (name *)
;
'params' ((name *) + 'and')
;
'consumes' nat?
;
\end{rail}
\begin{descr}
\item [$\CASE{(c~\vec x)}$] invokes a named local context $c\colon \vec x,
\vec \phi$, as provided by an appropriate proof method (such as $cases$ and
$induct$). The command ``$\CASE{(c~\vec x)}$'' abbreviates ``$\FIX{\vec
x}~\ASSUME{c}{\vec\phi}$''.
\item [$\isarkeyword{print_cases}$] prints all local contexts of the current
state, using Isar proof language notation. This is a diagnostic command;
$undo$ does not apply.
\item [$case_names~\vec c$] declares names for the local contexts of premises
of a theorem; $\vec c$ refers to the \emph{suffix} of the list of premises.
\item [$case_conclusion~c~\vec d$] declares names for the conclusions of a
named premise $c$; here $\vec d$ refers to the prefix of arguments of a
logical formula built by nesting a binary connective (e.g.\ $\lor$).
Note that proof methods such as $induct$ and $coinduct$ already provide a
default name for the conclusion as a whole. The need to name subformulas
only arises with cases that split into several sub-cases, as in common
co-induction rules.
\item [$params~\vec p@1 \dots \vec p@n$] renames the innermost parameters of
premises $1, \dots, n$ of some theorem. An empty list of names may be given
to skip positions, leaving the present parameters unchanged.
Note that the default usage of case rules does \emph{not} directly expose
parameters to the proof context.
\item [$consumes~n$] declares the number of ``major premises'' of a rule,
i.e.\ the number of facts to be consumed when it is applied by an
appropriate proof method. The default value of $consumes$ is $n = 1$, which
is appropriate for the usual kind of cases and induction rules for inductive
sets (cf.\ \S\ref{sec:hol-inductive}). Rules without any $consumes$
declaration given are treated as if $consumes~0$ had been specified.
Note that explicit $consumes$ declarations are only rarely needed; this is
already taken care of automatically by the higher-level $cases$, $induct$,
and $coinduct$ declarations.
\end{descr}
\subsubsection{Proof methods}
\indexisarmeth{cases}\indexisarmeth{induct}\indexisarmeth{coinduct}
\begin{matharray}{rcl}
cases & : & \isarmeth \\
induct & : & \isarmeth \\
coinduct & : & \isarmeth \\
\end{matharray}
The $cases$, $induct$, and $coinduct$ methods provide a uniform
interface to common proof techniques over datatypes, inductive
predicates (or sets), recursive functions etc. The corresponding
rules may be specified and instantiated in a casual manner.
Furthermore, these methods provide named local contexts that may be
invoked via the $\CASENAME$ proof command within the subsequent proof
text. This accommodates compact proof texts even when reasoning about
large specifications.
The $induct$ method also provides some additional infrastructure in order to
be applicable to structure statements (either using explicit meta-level
connectives, or including facts and parameters separately). This avoids
cumbersome encoding of ``strengthened'' inductive statements within the
object-logic.
\begin{rail}
'cases' (insts * 'and') rule?
;
'induct' (definsts * 'and') \\ arbitrary? taking? rule?
;
'coinduct' insts taking rule?
;
rule: ('type' | 'pred' | 'set') ':' (nameref +) | 'rule' ':' (thmref +)
;
definst: name ('==' | equiv) term | inst
;
definsts: ( definst *)
;
arbitrary: 'arbitrary' ':' ((term *) 'and' +)
;
taking: 'taking' ':' insts
;
\end{rail}
\begin{descr}
\item [$cases~insts~R$] applies method $rule$ with an appropriate case
distinction theorem, instantiated to the subjects $insts$. Symbolic case
names are bound according to the rule's local contexts.
The rule is determined as follows, according to the facts and arguments
passed to the $cases$ method:
\begin{matharray}{llll}
\Text{facts} & & \Text{arguments} & \Text{rule} \\\hline
& cases & & \Text{classical case split} \\
& cases & t & \Text{datatype exhaustion (type of $t$)} \\
\edrv A\; t & cases & \dots & \Text{inductive predicate/set elimination (of $A$)} \\
\dots & cases & \dots ~ R & \Text{explicit rule $R$} \\
\end{matharray}
Several instantiations may be given, referring to the \emph{suffix} of
premises of the case rule; within each premise, the \emph{prefix} of
variables is instantiated. In most situations, only a single term needs to
be specified; this refers to the first variable of the last premise (it is
usually the same for all cases).
\item [$induct~insts~R$] is analogous to the $cases$ method, but refers to
induction rules, which are determined as follows:
\begin{matharray}{llll}
\Text{facts} & & \Text{arguments} & \Text{rule} \\\hline
& induct & P ~ x ~ \dots & \Text{datatype induction (type of $x$)} \\
\edrv A\; x & induct & \dots & \Text{predicate/set induction (of $A$)} \\
\dots & induct & \dots ~ R & \Text{explicit rule $R$} \\
\end{matharray}
Several instantiations may be given, each referring to some part of
a mutual inductive definition or datatype --- only related partial
induction rules may be used together, though. Any of the lists of
terms $P, x, \dots$ refers to the \emph{suffix} of variables present
in the induction rule. This enables the writer to specify only
induction variables, or both predicates and variables, for example.
Instantiations may be definitional: equations $x \equiv t$ introduce local
definitions, which are inserted into the claim and discharged after applying
the induction rule. Equalities reappear in the inductive cases, but have
been transformed according to the induction principle being involved here.
In order to achieve practically useful induction hypotheses, some variables
occurring in $t$ need to be fixed (see below).
The optional ``$arbitrary\colon \vec x$'' specification generalizes
variables $\vec x$ of the original goal before applying induction. Thus
induction hypotheses may become sufficiently general to get the proof
through. Together with definitional instantiations, one may effectively
perform induction over expressions of a certain structure.
The optional ``$taking\colon \vec t$'' specification provides additional
instantiations of a prefix of pending variables in the rule. Such schematic
induction rules rarely occur in practice, though.
\item [$coinduct~inst~R$] is analogous to the $induct$ method, but refers to
coinduction rules, which are determined as follows:
\begin{matharray}{llll}
\Text{goal} & & \Text{arguments} & \Text{rule} \\\hline
& coinduct & x ~ \dots & \Text{type coinduction (type of $x$)} \\
A\; x & coinduct & \dots & \Text{predicate/set coinduction (of $A$)} \\
\dots & coinduct & \dots ~ R & \Text{explicit rule $R$} \\
\end{matharray}
Coinduction is the dual of induction. Induction essentially
eliminates $A\; x$ towards a generic result $P\; x$, while
coinduction introduces $A\; x$ starting with $B\; x$, for a suitable
``bisimulation'' $B$. The cases of a coinduct rule are typically
named after the predicates or sets being covered, while the
conclusions consist of several alternatives being named after the
individual destructor patterns.
The given instantiation refers to the \emph{suffix} of variables
occurring in the rule's major premise, or conclusion if unavailable.
An additional ``$taking: \vec t$'' specification may be required in
order to specify the bisimulation to be used in the coinduction
step.
\end{descr}
Above methods produce named local contexts, as determined by the instantiated
rule as given in the text. Beyond that, the $induct$ and $coinduct$ methods
guess further instantiations from the goal specification itself. Any
persisting unresolved schematic variables of the resulting rule will render
the the corresponding case invalid. The term binding
$\Var{case}$\indexisarvar{case} for the conclusion will be provided with each
case, provided that term is fully specified.
The $\isarkeyword{print_cases}$ command prints all named cases present in the
current proof state.
\medskip
Despite the additional infrastructure, both $cases$ and $coinduct$ merely
apply a certain rule, after instantiation, while conforming due to the usual
way of monotonic natural deduction: the context of a structured statement
$\All{\vec x} \vec\phi \Imp \dots$ reappears unchanged after the case split.
The $induct$ method is significantly different in this respect: the meta-level
structure is passed through the ``recursive'' course involved in the
induction. Thus the original statement is basically replaced by separate
copies, corresponding to the induction hypotheses and conclusion; the original
goal context is no longer available. Thus local assumptions, fixed parameters
and definitions effectively participate in the inductive rephrasing of the
original statement.
In induction proofs, local assumptions introduced by cases are split into two
different kinds: $hyps$ stemming from the rule and $prems$ from the goal
statement. This is reflected in the extracted cases accordingly, so invoking
``$\isarcmd{case}~c$'' will provide separate facts $c\mathord.hyps$ and
$c\mathord.prems$, as well as fact $c$ to hold the all-inclusive list.
\medskip
Facts presented to either method are consumed according to the number
of ``major premises'' of the rule involved, which is usually $0$ for
plain cases and induction rules of datatypes etc.\ and $1$ for rules
of inductive predicates or sets and the like. The remaining facts are
inserted into the goal verbatim before the actual $cases$, $induct$,
or $coinduct$ rule is applied.
\subsubsection{Declaring rules}
\indexisarcmd{print-induct-rules}\indexisaratt{cases}\indexisaratt{induct}\indexisaratt{coinduct}
\begin{matharray}{rcl}
\isarcmd{print_induct_rules}^* & : & \isarkeep{theory~|~proof} \\
cases & : & \isaratt \\
induct & : & \isaratt \\
coinduct & : & \isaratt \\
\end{matharray}
\begin{rail}
'cases' spec
;
'induct' spec
;
'coinduct' spec
;
spec: ('type' | 'pred' | 'set') ':' nameref
;
\end{rail}
\begin{descr}
\item [$\isarkeyword{print_induct_rules}$] prints cases and induct
rules for predicates (or sets) and types of the current context.
\item [$cases$, $induct$, and $coinduct$] (as attributes) augment the
corresponding context of rules for reasoning about (co)inductive
predicates (or sets) and types, using the corresponding methods of
the same name. Certain definitional packages of object-logics
usually declare emerging cases and induction rules as expected, so
users rarely need to intervene.
Manual rule declarations usually refer to the $case_names$ and
$params$ attributes to adjust names of cases and parameters of a
rule; the $consumes$ declaration is taken care of automatically:
$consumes~0$ is specified for ``type'' rules and $consumes~1$ for
``predicate'' / ``set'' rules.
\end{descr}
%%% Local Variables:
%%% mode: latex
%%% TeX-master: "isar-ref"
%%% End: