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
|
9272
|
14 |
\out{> Welcome to Isabelle/HOL (Isabelle99-1)}\medskip
|
7175
|
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}
|
9233
|
20 |
Note that any Isabelle/Isar command may be retracted by \texttt{undo}. See
|
|
21 |
the Isabelle/Isar Quick Reference (Appendix~{ap:refcard}) for a comprehensive
|
|
22 |
overview of available commands and other language elements.
|
7175
|
23 |
|
8508
|
24 |
|
8843
|
25 |
\subsection{Proof~General}
|
8508
|
26 |
|
|
27 |
Plain TTY-based interaction as above used to be quite feasible with
|
8547
|
28 |
traditional tactic based theorem proving, but developing Isar documents really
|
8508
|
29 |
demands some better user-interface support. David Aspinall's
|
|
30 |
\emph{Proof~General}\index{Proof General} environment
|
8547
|
31 |
\cite{proofgeneral,Aspinall:TACAS:2000} offers a generic Emacs interface for
|
|
32 |
interactive theorem provers that does all the cut-and-paste and
|
8508
|
33 |
forward-backward walk through the text in a very neat way. In Isabelle/Isar,
|
|
34 |
the current position within a partial proof document is equally important than
|
|
35 |
the actual proof state. Thus Proof~General provides the canonical working
|
|
36 |
environment for Isabelle/Isar, both for getting acquainted (e.g.\ by replaying
|
8547
|
37 |
existing Isar documents) and for production work.
|
7175
|
38 |
|
8843
|
39 |
|
|
40 |
\subsubsection{Proof~General as default Isabelle interface}
|
7167
|
41 |
|
7315
|
42 |
The easiest way to use Proof~General is to make it the default Isabelle user
|
8508
|
43 |
interface (see also \cite{isabelle-sys}). Just put something like this into
|
|
44 |
your Isabelle settings file:
|
7175
|
45 |
\begin{ttbox}
|
|
46 |
ISABELLE_INTERFACE=\$ISABELLE_HOME/contrib/ProofGeneral/isar/interface
|
8508
|
47 |
PROOFGENERAL_OPTIONS=""
|
7175
|
48 |
\end{ttbox}
|
|
49 |
You may have to change \texttt{\$ISABELLE_HOME/contrib/ProofGeneral} to the
|
7335
|
50 |
actual installation directory of Proof~General. From now on, the capital
|
7315
|
51 |
\texttt{Isabelle} executable refers to the \texttt{ProofGeneral/isar}
|
7981
|
52 |
interface.\footnote{There is also a \texttt{ProofGeneral/isa} interface, for
|
8508
|
53 |
classic Isabelle tactic scripts.}
|
7175
|
54 |
|
8547
|
55 |
The interface script provides several options, just pass \verb,-?, to see its
|
|
56 |
usage. Apart from the command line, the defaults for these options may be
|
|
57 |
overridden via the \texttt{PROOFGENERAL_OPTIONS} setting as well. For
|
8516
|
58 |
example, plain FSF Emacs (instead of the default XEmacs) may be configured in
|
|
59 |
Isabelle's settings via \texttt{PROOFGENERAL_OPTIONS="-p emacs"}.
|
7460
|
60 |
|
8516
|
61 |
Occasionally, the user's \verb,~/.emacs, file contains material that is
|
|
62 |
incompatible with the version of Emacs that Proof~General prefers. Then
|
8508
|
63 |
proper startup may be still achieved by using the \texttt{-u false} option.
|
8547
|
64 |
Also note that any Emacs lisp file called \texttt{proofgeneral-settings.el}
|
|
65 |
occurring in \texttt{\$ISABELLE_HOME/etc} or \texttt{\$ISABELLE_HOME_USER/etc}
|
|
66 |
is automatically loaded by the Proof~General interface script as well.
|
7981
|
67 |
|
|
68 |
\medskip
|
|
69 |
|
7175
|
70 |
With the proper Isabelle interface setup, Isar documents may now be edited by
|
|
71 |
visiting appropriate theory files, e.g.\
|
|
72 |
\begin{ttbox}
|
|
73 |
Isabelle \({\langle}isabellehome{\rangle}\)/src/HOL/Isar_examples/BasicLogic.thy
|
|
74 |
\end{ttbox}
|
7315
|
75 |
Users of XEmacs may note the tool bar for navigating forward and backward
|
8516
|
76 |
through the text. Consult the Proof~General documentation \cite{proofgeneral}
|
8547
|
77 |
for further basic command sequences, such as ``\texttt{C-c C-return}'' or
|
|
78 |
``\texttt{C-c u}''.
|
8508
|
79 |
|
8843
|
80 |
|
|
81 |
\subsubsection{The X-Symbol package}
|
8508
|
82 |
|
|
83 |
Proof~General also supports the Emacs X-Symbol package \cite{x-symbol}, which
|
|
84 |
provides a nice way to get proper mathematical symbols displayed on screen.
|
|
85 |
Just pass option \texttt{-x true} to the Isabelle interface script, or check
|
8516
|
86 |
the appropriate menu setting by hand. In any case, the X-Symbol package must
|
|
87 |
have been properly installed already.
|
|
88 |
|
8843
|
89 |
Contrary to what you may expect from the documentation of X-Symbol, the
|
|
90 |
package is very easy to install for individual users and configures itself
|
|
91 |
automatically. Simply download the ``binary'' package file, and do something
|
|
92 |
like this to install it in your home directory:
|
|
93 |
\begin{ttbox}
|
|
94 |
mkdir -p ~/.xemacs
|
|
95 |
cd ~/.xemacs
|
|
96 |
tar xzf .../x-symbol-pkg.tar.gz
|
|
97 |
\end{ttbox}
|
|
98 |
|
|
99 |
\medskip
|
|
100 |
|
|
101 |
Using proper mathematical symbols in Isabelle theories can be very convenient
|
|
102 |
for readability of large formulas. On the other hand, the plain ASCII sources
|
|
103 |
easily become somewhat unintelligible. For example, $\forall$ will appear as
|
|
104 |
\verb,\\<forall>, according the default set of Isabelle symbols.
|
|
105 |
Nevertheless, the Isabelle document preparation system (see
|
|
106 |
\S\ref{sec:document-prep}) will be happy to print non-ASCII symbols properly.
|
|
107 |
It is even possible to invent additional notation beyond the display
|
|
108 |
capabilities of XEmacs and X-Symbol.
|
7175
|
109 |
|
7981
|
110 |
|
|
111 |
\section{Isabelle/Isar theories}
|
|
112 |
|
8547
|
113 |
Isabelle/Isar offers the following main improvements over classic Isabelle.
|
7981
|
114 |
\begin{enumerate}
|
|
115 |
\item A new \emph{theory format}, occasionally referred to as ``new-style
|
|
116 |
theories'', supporting interactive development and unlimited undo operation.
|
|
117 |
\item A \emph{formal proof document language} designed to support intelligible
|
|
118 |
semi-automated reasoning. Instead of putting together unreadable tactic
|
|
119 |
scripts, the author is enabled to express the reasoning in way that is close
|
8508
|
120 |
to usual mathematical practice.
|
8547
|
121 |
\item A simple document preparation system, for typesetting formal
|
|
122 |
developments together with informal text. The resulting hyper-linked PDF
|
|
123 |
documents are equally well suited for WWW presentation and as printed
|
|
124 |
copies.
|
7981
|
125 |
\end{enumerate}
|
|
126 |
|
|
127 |
The Isar proof language is embedded into the new theory format as a proper
|
|
128 |
sub-language. Proof mode is entered by stating some $\THEOREMNAME$ or
|
|
129 |
$\LEMMANAME$ at the theory level, and left again with the final conclusion
|
|
130 |
(e.g.\ via $\QEDNAME$). A few theory extension mechanisms require proof as
|
8547
|
131 |
well, such as HOL's $\isarkeyword{typedef}$ which demands non-emptiness of the
|
|
132 |
representing sets.
|
7460
|
133 |
|
7981
|
134 |
New-style theory files may still be associated with separate ML files
|
|
135 |
consisting of plain old tactic scripts. There is no longer any ML binding
|
|
136 |
generated for the theory and theorems, though. ML functions \texttt{theory},
|
|
137 |
\texttt{thm}, and \texttt{thms} retrieve this information \cite{isabelle-ref}.
|
|
138 |
Nevertheless, migration between classic Isabelle and Isabelle/Isar is
|
|
139 |
relatively easy. Thus users may start to benefit from interactive theory
|
8547
|
140 |
development and document preparation, even before they have any idea of the
|
|
141 |
Isar proof language at all.
|
7981
|
142 |
|
|
143 |
\begin{warn}
|
8547
|
144 |
Currently, Proof~General does \emph{not} support mixed interactive
|
7981
|
145 |
development of classic Isabelle theory files or tactic scripts, together
|
|
146 |
with Isar documents. The ``\texttt{isa}'' and ``\texttt{isar}'' versions of
|
|
147 |
Proof~General are handled as two different theorem proving systems, only one
|
|
148 |
of these may be active at the same time.
|
|
149 |
\end{warn}
|
|
150 |
|
|
151 |
Porting of existing tactic scripts is best done by running two separate
|
|
152 |
Proof~General sessions, one for replaying the old script and the other for the
|
|
153 |
emerging Isabelle/Isar document.
|
|
154 |
|
7167
|
155 |
|
8843
|
156 |
\subsection{Document preparation}\label{sec:document-prep}
|
8684
|
157 |
|
|
158 |
Isabelle/Isar provides a simple document preparation system based on current
|
|
159 |
(PDF) {\LaTeX} technology, with full support of hyper-links (both local
|
|
160 |
references and URLs), bookmarks, thumbnails etc. Thus the results are equally
|
|
161 |
well suited for WWW browsing and as printed copies.
|
|
162 |
|
|
163 |
\medskip
|
|
164 |
|
|
165 |
Isabelle generates {\LaTeX} output as part of the run of a \emph{logic
|
|
166 |
session} (see also \cite{isabelle-sys}). Getting started with a working
|
|
167 |
configuration for common situations is quite easy by using the Isabelle
|
|
168 |
\texttt{mkdir} and \texttt{make} tools. Just invoke
|
|
169 |
\begin{ttbox}
|
|
170 |
isatool mkdir -d Foo
|
|
171 |
\end{ttbox}
|
|
172 |
to setup a separate directory for session \texttt{Foo}.\footnote{It is safe to
|
|
173 |
experiment, since \texttt{isatool mkdir} never overwrites existing files.}
|
|
174 |
Ensure that \texttt{Foo/ROOT.ML} loads all theories required for this session.
|
|
175 |
Furthermore \texttt{Foo/document/root.tex} should include any special {\LaTeX}
|
|
176 |
macro packages required for your document (the default is usually sufficient
|
|
177 |
as a start).
|
|
178 |
|
|
179 |
The session is controlled by a separate \texttt{IsaMakefile} (with very crude
|
|
180 |
source dependencies only by default). This file is located one level up from
|
|
181 |
the \texttt{Foo} directory location. At that point just invoke
|
|
182 |
\begin{ttbox}
|
|
183 |
isatool make Foo
|
|
184 |
\end{ttbox}
|
|
185 |
to run the \texttt{Foo} session, with browser information and document
|
|
186 |
preparation enabled. Unless any errors are reported by Isabelle or {\LaTeX},
|
|
187 |
the output will appear inside the directory indicated by \texttt{isatool
|
|
188 |
getenv ISABELLE_BROWSER_INFO}, with the logical session prefix added (e.g.\
|
|
189 |
\texttt{HOL/Foo}). Note that the \texttt{index.html} located there provides a
|
|
190 |
link to the finished {\LaTeX} document, too.
|
|
191 |
|
|
192 |
Note that this really is batch processing --- better let Isabelle check your
|
|
193 |
theory and proof developments beforehand in interactive mode.
|
|
194 |
|
|
195 |
\medskip
|
|
196 |
|
|
197 |
You may also consider to tune the \texttt{usedir} options in
|
|
198 |
\texttt{IsaMakefile}, for example to change the output format from
|
|
199 |
\texttt{dvi} to \texttt{pdf}, or activate the \texttt{-D document} option in
|
|
200 |
order to preserve a copy of the generated {\LaTeX} sources. The latter
|
|
201 |
feature is very useful for debugging {\LaTeX} errors, while avoiding repeated
|
|
202 |
runs of Isabelle.
|
|
203 |
|
|
204 |
\medskip
|
|
205 |
|
|
206 |
See \emph{The Isabelle System Manual} \cite{isabelle-sys} for further details
|
|
207 |
on Isabelle logic sessions and theory presentation.
|
|
208 |
|
|
209 |
|
|
210 |
\subsection{How to write Isar proofs anyway?}
|
7167
|
211 |
|
7297
|
212 |
This is one of the key questions, of course. Isar offers a rather different
|
|
213 |
approach to formal proof documents than plain old tactic scripts. Experienced
|
|
214 |
users of existing interactive theorem proving systems may have to learn
|
7895
|
215 |
thinking differently in order to make effective use of Isabelle/Isar. On the
|
7297
|
216 |
other hand, Isabelle/Isar comes much closer to existing mathematical practice
|
|
217 |
of formal proof, so users with less experience in old-style tactical proving,
|
7895
|
218 |
but a good understanding of mathematical proof, might cope with Isar even
|
7981
|
219 |
better. See also \cite{Wenzel:1999:TPHOL} for further background information
|
|
220 |
on Isar.
|
8547
|
221 |
%FIXME cite [HahnBanach-in-Isar]
|
7297
|
222 |
|
7981
|
223 |
\medskip This really is a \emph{reference manual}. Nevertheless, we will also
|
|
224 |
give some clues of how the concepts introduced here may be put into practice.
|
|
225 |
Appendix~\ref{ap:refcard} provides a quick reference card of the most common
|
|
226 |
Isabelle/Isar language elements. There are several examples distributed with
|
8516
|
227 |
Isabelle, and available via the Isabelle WWW library as well as the
|
|
228 |
Isabelle/Isar page:
|
7836
|
229 |
\begin{center}\small
|
|
230 |
\begin{tabular}{l}
|
|
231 |
\url{http://www.cl.cam.ac.uk/Research/HVG/Isabelle/library/} \\
|
8516
|
232 |
\url{http://isabelle.in.tum.de/library/} \\[1ex]
|
8508
|
233 |
\url{http://isabelle.in.tum.de/Isar/} \\
|
7836
|
234 |
\end{tabular}
|
|
235 |
\end{center}
|
|
236 |
|
7987
|
237 |
See \texttt{HOL/Isar_examples} for a collection of introductory examples, and
|
8547
|
238 |
\texttt{HOL/HOL-Real/HahnBanach} for a big mathematics application. Apart
|
|
239 |
from plain HTML sources, these sessions also provide actual documents (in
|
|
240 |
PDF).
|
|
241 |
|
7167
|
242 |
|
7046
|
243 |
%%% Local Variables:
|
|
244 |
%%% mode: latex
|
|
245 |
%%% TeX-master: "isar-ref"
|
|
246 |
%%% End:
|