2
|
1 |
theory Intro
|
|
2 |
imports Main
|
|
3 |
|
|
4 |
begin
|
|
5 |
|
|
6 |
chapter {* Introduction *}
|
|
7 |
|
|
8 |
text {*
|
5
|
9 |
The purpose of this cookbook is to guide the reader through the
|
11
|
10 |
first steps of Isabelle programming, and to provide recipes for
|
2
|
11 |
solving common problems.
|
|
12 |
*}
|
|
13 |
|
|
14 |
section {* Intended Audience and Prior Knowledge *}
|
|
15 |
|
|
16 |
text {*
|
5
|
17 |
This cookbook targets an audience who already knows how to use Isabelle
|
6
|
18 |
for writing theories and proofs. We also assume that readers are
|
50
|
19 |
familiar with Standard ML (SML), the programming language in which
|
11
|
20 |
most of Isabelle is implemented. If you are unfamiliar with either of
|
2
|
21 |
these two subjects, you should first work through the Isabelle/HOL
|
50
|
22 |
tutorial \cite{isa-tutorial} or Paulson's book on SML
|
2
|
23 |
\cite{paulson-ml2}.
|
|
24 |
|
|
25 |
*}
|
|
26 |
|
5
|
27 |
section {* Existing Documentation *}
|
2
|
28 |
|
|
29 |
text {*
|
|
30 |
|
43
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
31 |
The following documentation about Isabelle programming already exists (and is
|
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
32 |
part of the distribution of Isabelle):
|
2
|
33 |
|
|
34 |
\begin{description}
|
5
|
35 |
\item[The Implementation Manual] describes Isabelle
|
44
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
36 |
from a high-level perspective, documenting both the underlying
|
6
|
37 |
concepts and some of the interfaces.
|
2
|
38 |
|
5
|
39 |
\item[The Isabelle Reference Manual] is an older document that used
|
44
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
40 |
to be the main reference of Isabelle at a time when all proof scripts
|
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
41 |
were written on the ML level. Many parts of this manual are outdated
|
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
42 |
now, but some parts, particularly the chapters on tactics, are still
|
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
43 |
useful.
|
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
44 |
|
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
45 |
\item[The Isar Reference Manual] is also an older document that provides
|
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
46 |
material about Isar and its implementation. Some material in it
|
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
47 |
is still useful.
|
5
|
48 |
\end{description}
|
2
|
49 |
|
12
|
50 |
Then of course there is:
|
5
|
51 |
|
|
52 |
\begin{description}
|
2
|
53 |
\item[The code] is of course the ultimate reference for how
|
|
54 |
things really work. Therefore you should not hesitate to look at the
|
|
55 |
way things are actually implemented. More importantly, it is often
|
|
56 |
good to look at code that does similar things as you want to do, to
|
|
57 |
learn from other people's code.
|
|
58 |
\end{description}
|
|
59 |
|
50
|
60 |
The Cookbook is written in such a way that the code examples in it are
|
|
61 |
synchronised with fairly recent versions of the code.
|
|
62 |
|
2
|
63 |
*}
|
|
64 |
|
|
65 |
|
|
66 |
end |