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