| author | Christian Urban <urbanc@in.tum.de> | 
| Tue, 24 Feb 2009 11:15:41 +0000 | |
| changeset 134 | 328370b75c33 | 
| parent 126 | fcc0e6e54dca | 
| child 137 | a9685909944d | 
| permissions | -rw-r--r-- | 
| 
2
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1  | 
theory Intro  | 
| 75 | 2  | 
imports Base  | 
| 
2
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
4  | 
begin  | 
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
5  | 
|
| 
64
 
9a6e5e0c4906
deleted old files and added code to give a special tag to the command ML
 
Christian Urban <urbanc@in.tum.de> 
parents: 
60 
diff
changeset
 | 
6  | 
|
| 
2
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
7  | 
chapter {* Introduction *}
 | 
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
9  | 
text {*
 | 
| 89 | 10  | 
If your next project requires you to program on the ML-level of Isabelle,  | 
| 
106
 
bdd82350cf22
renamed in the pdf all instances of cookbook to tutorial (in order to sound more serious)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
102 
diff
changeset
 | 
11  | 
then this tutorial is for you. It will guide you through the first steps of  | 
| 89 | 12  | 
Isabelle programming, and also explain tricks of the trade. The best way to  | 
13  | 
get to know the ML-level of Isabelle is by experimenting with the many code  | 
|
| 
106
 
bdd82350cf22
renamed in the pdf all instances of cookbook to tutorial (in order to sound more serious)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
102 
diff
changeset
 | 
14  | 
examples included in the tutorial. The code is as far as possible checked  | 
| 89 | 15  | 
against recent versions of Isabelle. If something does not work, then  | 
16  | 
please let us know. If you have comments, criticism or like to add to the  | 
|
| 121 | 17  | 
tutorial, feel free---you are most welcome! The tutorial is meant to be  | 
| 126 | 18  | 
gentle and comprehensive. To achieve this we need your feedback.  | 
| 
2
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
19  | 
*}  | 
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
20  | 
|
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
21  | 
section {* Intended Audience and Prior Knowledge *}
 | 
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
22  | 
|
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
23  | 
text {* 
 | 
| 
106
 
bdd82350cf22
renamed in the pdf all instances of cookbook to tutorial (in order to sound more serious)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
102 
diff
changeset
 | 
24  | 
This tutorial targets readers who already know how to use Isabelle for  | 
| 75 | 25  | 
writing theories and proofs. We also assume that readers are familiar with  | 
26  | 
the functional programming language ML, the language in which most of  | 
|
27  | 
Isabelle is implemented. If you are unfamiliar with either of these two  | 
|
28  | 
subjects, you should first work through the Isabelle/HOL tutorial  | 
|
29  | 
  \cite{isa-tutorial} or Paulson's book on ML \cite{paulson-ml2}.
 | 
|
| 
2
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
31  | 
*}  | 
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
32  | 
|
| 
5
 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2 
diff
changeset
 | 
33  | 
section {* Existing Documentation *}
 | 
| 
2
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
34  | 
|
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
35  | 
text {*
 | 
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
36  | 
|
| 
43
 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 
Christian Urban <urbanc@in.tum.de> 
parents: 
42 
diff
changeset
 | 
37  | 
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> 
parents: 
42 
diff
changeset
 | 
38  | 
part of the distribution of Isabelle):  | 
| 
2
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
39  | 
|
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
40  | 
  \begin{description}
 | 
| 
5
 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2 
diff
changeset
 | 
41  | 
\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> 
parents: 
43 
diff
changeset
 | 
42  | 
from a high-level perspective, documenting both the underlying  | 
| 6 | 43  | 
concepts and some of the interfaces.  | 
| 
2
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
44  | 
|
| 
5
 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2 
diff
changeset
 | 
45  | 
\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> 
parents: 
43 
diff
changeset
 | 
46  | 
to be the main reference of Isabelle at a time when all proof scripts  | 
| 89 | 47  | 
were written on the ML-level. Many parts of this manual are outdated  | 
| 
44
 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
48  | 
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> 
parents: 
43 
diff
changeset
 | 
49  | 
useful.  | 
| 
 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
50  | 
|
| 
 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
51  | 
\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> 
parents: 
43 
diff
changeset
 | 
52  | 
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> 
parents: 
43 
diff
changeset
 | 
53  | 
is still useful.  | 
| 
5
 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2 
diff
changeset
 | 
54  | 
  \end{description}
 | 
| 
2
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
55  | 
|
| 
12
 
2f1736cb8f26
various changes by Alex and Christian
 
Christian Urban <urbanc@in.tum.de> 
parents: 
11 
diff
changeset
 | 
56  | 
Then of course there is:  | 
| 
5
 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2 
diff
changeset
 | 
57  | 
|
| 
 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2 
diff
changeset
 | 
58  | 
  \begin{description}
 | 
| 
2
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
59  | 
\item[The code] is of course the ultimate reference for how  | 
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
60  | 
things really work. Therefore you should not hesitate to look at the  | 
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
61  | 
way things are actually implemented. More importantly, it is often  | 
| 89 | 62  | 
good to look at code that does similar things as you want to do and  | 
63  | 
to learn from other people's code.  | 
|
| 
2
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
64  | 
  \end{description}
 | 
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
65  | 
|
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
66  | 
*}  | 
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
67  | 
|
| 81 | 68  | 
section {* Typographic Conventions *}
 | 
| 
68
 
e7519207c2b7
added more to the "new command section" and tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
66 
diff
changeset
 | 
69  | 
|
| 
 
e7519207c2b7
added more to the "new command section" and tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
66 
diff
changeset
 | 
70  | 
text {*
 | 
| 75 | 71  | 
|
| 
106
 
bdd82350cf22
renamed in the pdf all instances of cookbook to tutorial (in order to sound more serious)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
102 
diff
changeset
 | 
72  | 
All ML-code in this tutorial is typeset in highlighted boxes, like the following  | 
| 89 | 73  | 
ML-expression:  | 
| 75 | 74  | 
|
75  | 
  \begin{isabelle}
 | 
|
76  | 
  \begin{graybox}
 | 
|
| 
85
 
b02904872d6b
better handling of {* and *}
 
Christian Urban <urbanc@in.tum.de> 
parents: 
84 
diff
changeset
 | 
77  | 
  \isacommand{ML}~@{text "\<verbopen>"}\isanewline
 | 
| 75 | 78  | 
  \hspace{5mm}@{ML "3 + 4"}\isanewline
 | 
| 
85
 
b02904872d6b
better handling of {* and *}
 
Christian Urban <urbanc@in.tum.de> 
parents: 
84 
diff
changeset
 | 
79  | 
  @{text "\<verbclose>"}
 | 
| 75 | 80  | 
  \end{graybox}
 | 
81  | 
  \end{isabelle}
 | 
|
82  | 
||
| 89 | 83  | 
These boxes corresponds to how code can be processed inside the interactive  | 
84  | 
environment of Isabelle. It is therefore easy to experiment with what is  | 
|
85  | 
displayed. However, for better readability we will drop the enclosing  | 
|
86  | 
  \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"} and just write:
 | 
|
| 75 | 87  | 
|
| 81 | 88  | 
|
| 75 | 89  | 
  @{ML [display,gray] "3 + 4"}
 | 
90  | 
||
| 89 | 91  | 
Whenever appropriate we also show the response the code  | 
| 81 | 92  | 
generates when evaluated. This response is prefixed with a  | 
| 
108
 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
107 
diff
changeset
 | 
93  | 
  @{text [quotes] ">"}, like:
 | 
| 75 | 94  | 
|
95  | 
  @{ML_response [display,gray] "3 + 4" "7"}
 | 
|
96  | 
||
| 122 | 97  | 
The user-level commands of Isabelle (i.e.~the non-ML code) are written  | 
98  | 
  in bold, for example \isacommand{lemma}, \isacommand{apply},
 | 
|
99  | 
  \isacommand{foobar} and so on.  We use @{text "$"} to indicate that a
 | 
|
100  | 
command needs to be run in a Unix-shell, for example:  | 
|
| 75 | 101  | 
|
102  | 
  @{text [display] "$ ls -la"}
 | 
|
103  | 
||
| 89 | 104  | 
Pointers to further information and Isabelle files are typeset in  | 
105  | 
italic and highlighted as follows:  | 
|
| 75 | 106  | 
|
107  | 
  \begin{readmore}
 | 
|
| 
102
 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 
Christian Urban <urbanc@in.tum.de> 
parents: 
101 
diff
changeset
 | 
108  | 
Further information or pointers to files.  | 
| 75 | 109  | 
  \end{readmore}
 | 
110  | 
||
| 
68
 
e7519207c2b7
added more to the "new command section" and tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
66 
diff
changeset
 | 
111  | 
*}  | 
| 
 
e7519207c2b7
added more to the "new command section" and tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
66 
diff
changeset
 | 
112  | 
|
| 
119
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
113  | 
section {* Acknowledgements *}
 | 
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
114  | 
|
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
115  | 
text {*
 | 
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
116  | 
Financial support for this tutorial was provided by the German  | 
| 122 | 117  | 
Research Council (DFG) under grant number URB 165/5-1. The following  | 
118  | 
people contributed:  | 
|
| 
119
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
119  | 
|
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
120  | 
  \begin{itemize}
 | 
| 122 | 121  | 
  \item {\bf Stefan Berghofer} wrote nearly all of the ML-code of the
 | 
122  | 
  \simpleinductive-package and the code for the @{text
 | 
|
123  | 
"chunk"}-antiquotation. He also wrote the first version of the chapter  | 
|
124  | 
  describing the package and has been helpful \emph{beyond measure} with
 | 
|
125  | 
answering questions about Isabelle.  | 
|
| 
119
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
126  | 
|
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
127  | 
  \item {\bf Sascha Böhme} contributed the recipes in \ref{rec:timeout}, 
 | 
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
128  | 
  \ref{rec:config}, \ref{rec:storingdata}, \ref{rec:external} and \ref{rec:oracle}.
 | 
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
129  | 
|
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
130  | 
  \item {\bf Jeremy Dawson} wrote the first version of the chapter
 | 
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
131  | 
about parsing.  | 
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
132  | 
|
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
133  | 
  \item {\bf Alexander Krauss} wrote the first version of the ``first-steps'' 
 | 
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
134  | 
  chapter and also contributed recipe \ref{rec:named}.
 | 
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
135  | 
  \end{itemize}
 | 
| 
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
136  | 
|
| 121 | 137  | 
Please let me know of any omissions. Responsibility for any remaining  | 
138  | 
errors lies with me.  | 
|
| 
119
 
4536782969fa
added an acknowledgement section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
108 
diff
changeset
 | 
139  | 
*}  | 
| 
2
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
140  | 
|
| 
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
141  | 
end  |