7046
|
1 |
|
|
2 |
\chapter{Introduction}
|
|
3 |
|
7167
|
4 |
\section{Quick start}
|
|
5 |
|
8508
|
6 |
\subsection{Terminal sessions}
|
|
7 |
|
7175
|
8 |
Isar is already part of Isabelle (as of version Isabelle99, or later). The
|
|
9 |
\texttt{isabelle} binary provides option \texttt{-I} to run the Isar
|
|
10 |
interaction loop at startup, rather than the plain ML top-level. Thus the
|
8547
|
11 |
quickest way to do \emph{anything} with Isabelle/Isar is as follows:
|
7175
|
12 |
\begin{ttbox}
|
|
13 |
isabelle -I HOL\medskip
|
|
14 |
\out{> Welcome to Isabelle/HOL (Isabelle99)}\medskip
|
|
15 |
theory Foo = Main:
|
7297
|
16 |
constdefs foo :: nat "foo == 1";
|
|
17 |
lemma "0 < foo" by (simp add: foo_def);
|
7175
|
18 |
end
|
|
19 |
\end{ttbox}
|
8547
|
20 |
Note that any Isabelle/Isar command may be retracted by \texttt{undo}. The
|
7981
|
21 |
\texttt{help} command prints a list of available language elements.
|
7175
|
22 |
|
8508
|
23 |
|
|
24 |
\subsection{The Proof~General interface}
|
|
25 |
|
|
26 |
Plain TTY-based interaction as above used to be quite feasible with
|
8547
|
27 |
traditional tactic based theorem proving, but developing Isar documents really
|
8508
|
28 |
demands some better user-interface support. David Aspinall's
|
|
29 |
\emph{Proof~General}\index{Proof General} environment
|
8547
|
30 |
\cite{proofgeneral,Aspinall:TACAS:2000} offers a generic Emacs interface for
|
|
31 |
interactive theorem provers that does all the cut-and-paste and
|
8508
|
32 |
forward-backward walk through the text in a very neat way. In Isabelle/Isar,
|
|
33 |
the current position within a partial proof document is equally important than
|
|
34 |
the actual proof state. Thus Proof~General provides the canonical working
|
|
35 |
environment for Isabelle/Isar, both for getting acquainted (e.g.\ by replaying
|
8547
|
36 |
existing Isar documents) and for production work.
|
7175
|
37 |
|
|
38 |
\medskip
|
7167
|
39 |
|
7315
|
40 |
The easiest way to use Proof~General is to make it the default Isabelle user
|
8508
|
41 |
interface (see also \cite{isabelle-sys}). Just put something like this into
|
|
42 |
your Isabelle settings file:
|
7175
|
43 |
\begin{ttbox}
|
|
44 |
ISABELLE_INTERFACE=\$ISABELLE_HOME/contrib/ProofGeneral/isar/interface
|
8508
|
45 |
PROOFGENERAL_OPTIONS=""
|
7175
|
46 |
\end{ttbox}
|
|
47 |
You may have to change \texttt{\$ISABELLE_HOME/contrib/ProofGeneral} to the
|
7335
|
48 |
actual installation directory of Proof~General. From now on, the capital
|
7315
|
49 |
\texttt{Isabelle} executable refers to the \texttt{ProofGeneral/isar}
|
7981
|
50 |
interface.\footnote{There is also a \texttt{ProofGeneral/isa} interface, for
|
8508
|
51 |
classic Isabelle tactic scripts.}
|
7175
|
52 |
|
8547
|
53 |
The interface script provides several options, just pass \verb,-?, to see its
|
|
54 |
usage. Apart from the command line, the defaults for these options may be
|
|
55 |
overridden via the \texttt{PROOFGENERAL_OPTIONS} setting as well. For
|
8516
|
56 |
example, plain FSF Emacs (instead of the default XEmacs) may be configured in
|
|
57 |
Isabelle's settings via \texttt{PROOFGENERAL_OPTIONS="-p emacs"}.
|
7460
|
58 |
|
8516
|
59 |
Occasionally, the user's \verb,~/.emacs, file contains material that is
|
|
60 |
incompatible with the version of Emacs that Proof~General prefers. Then
|
8508
|
61 |
proper startup may be still achieved by using the \texttt{-u false} option.
|
8547
|
62 |
Also note that any Emacs lisp file called \texttt{proofgeneral-settings.el}
|
|
63 |
occurring in \texttt{\$ISABELLE_HOME/etc} or \texttt{\$ISABELLE_HOME_USER/etc}
|
|
64 |
is automatically loaded by the Proof~General interface script as well.
|
7981
|
65 |
|
|
66 |
\medskip
|
|
67 |
|
7175
|
68 |
With the proper Isabelle interface setup, Isar documents may now be edited by
|
|
69 |
visiting appropriate theory files, e.g.\
|
|
70 |
\begin{ttbox}
|
|
71 |
Isabelle \({\langle}isabellehome{\rangle}\)/src/HOL/Isar_examples/BasicLogic.thy
|
|
72 |
\end{ttbox}
|
7315
|
73 |
Users of XEmacs may note the tool bar for navigating forward and backward
|
8516
|
74 |
through the text. Consult the Proof~General documentation \cite{proofgeneral}
|
8547
|
75 |
for further basic command sequences, such as ``\texttt{C-c C-return}'' or
|
|
76 |
``\texttt{C-c u}''.
|
8508
|
77 |
|
|
78 |
\medskip
|
|
79 |
|
|
80 |
Proof~General also supports the Emacs X-Symbol package \cite{x-symbol}, which
|
|
81 |
provides a nice way to get proper mathematical symbols displayed on screen.
|
|
82 |
Just pass option \texttt{-x true} to the Isabelle interface script, or check
|
8516
|
83 |
the appropriate menu setting by hand. In any case, the X-Symbol package must
|
|
84 |
have been properly installed already.
|
|
85 |
|
|
86 |
Note that using actual mathematical symbols in the text easily makes the ASCII
|
|
87 |
sources hard to read. For example, $\forall$ will appear as \verb,\\<forall>,
|
|
88 |
according the default set of Isabelle symbols. On the other hand, the
|
|
89 |
Isabelle document preparation system \cite{isabelle-sys} will be happy to
|
|
90 |
print non-ASCII symbols properly. It is even possible to invent additional
|
8547
|
91 |
notation beyond the display capabilities of XEmacs and X-Symbol.
|
7175
|
92 |
|
7981
|
93 |
|
|
94 |
\section{Isabelle/Isar theories}
|
|
95 |
|
8547
|
96 |
Isabelle/Isar offers the following main improvements over classic Isabelle.
|
7981
|
97 |
\begin{enumerate}
|
|
98 |
\item A new \emph{theory format}, occasionally referred to as ``new-style
|
|
99 |
theories'', supporting interactive development and unlimited undo operation.
|
|
100 |
\item A \emph{formal proof document language} designed to support intelligible
|
|
101 |
semi-automated reasoning. Instead of putting together unreadable tactic
|
|
102 |
scripts, the author is enabled to express the reasoning in way that is close
|
8508
|
103 |
to usual mathematical practice.
|
8547
|
104 |
\item A simple document preparation system, for typesetting formal
|
|
105 |
developments together with informal text. The resulting hyper-linked PDF
|
|
106 |
documents are equally well suited for WWW presentation and as printed
|
|
107 |
copies.
|
7981
|
108 |
\end{enumerate}
|
|
109 |
|
|
110 |
The Isar proof language is embedded into the new theory format as a proper
|
|
111 |
sub-language. Proof mode is entered by stating some $\THEOREMNAME$ or
|
|
112 |
$\LEMMANAME$ at the theory level, and left again with the final conclusion
|
|
113 |
(e.g.\ via $\QEDNAME$). A few theory extension mechanisms require proof as
|
8547
|
114 |
well, such as HOL's $\isarkeyword{typedef}$ which demands non-emptiness of the
|
|
115 |
representing sets.
|
7460
|
116 |
|
7981
|
117 |
New-style theory files may still be associated with separate ML files
|
|
118 |
consisting of plain old tactic scripts. There is no longer any ML binding
|
|
119 |
generated for the theory and theorems, though. ML functions \texttt{theory},
|
|
120 |
\texttt{thm}, and \texttt{thms} retrieve this information \cite{isabelle-ref}.
|
|
121 |
Nevertheless, migration between classic Isabelle and Isabelle/Isar is
|
|
122 |
relatively easy. Thus users may start to benefit from interactive theory
|
8547
|
123 |
development and document preparation, even before they have any idea of the
|
|
124 |
Isar proof language at all.
|
7981
|
125 |
|
|
126 |
\begin{warn}
|
8547
|
127 |
Currently, Proof~General does \emph{not} support mixed interactive
|
7981
|
128 |
development of classic Isabelle theory files or tactic scripts, together
|
|
129 |
with Isar documents. The ``\texttt{isa}'' and ``\texttt{isar}'' versions of
|
|
130 |
Proof~General are handled as two different theorem proving systems, only one
|
|
131 |
of these may be active at the same time.
|
|
132 |
\end{warn}
|
|
133 |
|
|
134 |
Porting of existing tactic scripts is best done by running two separate
|
|
135 |
Proof~General sessions, one for replaying the old script and the other for the
|
|
136 |
emerging Isabelle/Isar document.
|
|
137 |
|
7167
|
138 |
|
|
139 |
\section{How to write Isar proofs anyway?}
|
|
140 |
|
7297
|
141 |
This is one of the key questions, of course. Isar offers a rather different
|
|
142 |
approach to formal proof documents than plain old tactic scripts. Experienced
|
|
143 |
users of existing interactive theorem proving systems may have to learn
|
7895
|
144 |
thinking differently in order to make effective use of Isabelle/Isar. On the
|
7297
|
145 |
other hand, Isabelle/Isar comes much closer to existing mathematical practice
|
|
146 |
of formal proof, so users with less experience in old-style tactical proving,
|
7895
|
147 |
but a good understanding of mathematical proof, might cope with Isar even
|
7981
|
148 |
better. See also \cite{Wenzel:1999:TPHOL} for further background information
|
|
149 |
on Isar.
|
8547
|
150 |
%FIXME cite [HahnBanach-in-Isar]
|
7297
|
151 |
|
7981
|
152 |
\medskip This really is a \emph{reference manual}. Nevertheless, we will also
|
|
153 |
give some clues of how the concepts introduced here may be put into practice.
|
|
154 |
Appendix~\ref{ap:refcard} provides a quick reference card of the most common
|
|
155 |
Isabelle/Isar language elements. There are several examples distributed with
|
8516
|
156 |
Isabelle, and available via the Isabelle WWW library as well as the
|
|
157 |
Isabelle/Isar page:
|
7836
|
158 |
\begin{center}\small
|
|
159 |
\begin{tabular}{l}
|
|
160 |
\url{http://www.cl.cam.ac.uk/Research/HVG/Isabelle/library/} \\
|
8516
|
161 |
\url{http://isabelle.in.tum.de/library/} \\[1ex]
|
8508
|
162 |
\url{http://isabelle.in.tum.de/Isar/} \\
|
7836
|
163 |
\end{tabular}
|
|
164 |
\end{center}
|
|
165 |
|
7987
|
166 |
See \texttt{HOL/Isar_examples} for a collection of introductory examples, and
|
8547
|
167 |
\texttt{HOL/HOL-Real/HahnBanach} for a big mathematics application. Apart
|
|
168 |
from plain HTML sources, these sessions also provide actual documents (in
|
|
169 |
PDF).
|
|
170 |
|
7167
|
171 |
|
7046
|
172 |
%%% Local Variables:
|
|
173 |
%%% mode: latex
|
|
174 |
%%% TeX-master: "isar-ref"
|
|
175 |
%%% End:
|