| author | Christian Urban <urbanc@in.tum.de> | 
| Thu, 29 Jan 2009 17:09:56 +0000 | |
| changeset 91 | 667a0943c40b | 
| parent 89 | fee4942c4770 | 
| child 101 | 123401a5c8e9 | 
| 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: 
60diff
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, | 
| 11 | then this Cookbook is for you. It will guide you through the first steps of | |
| 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 | |
| 14 | examples included in the Cookbook. The code is as far as possible checked | |
| 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 | |
| 17 | Cookbook, feel free---you are most welcome! | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 18 | *} | 
| 
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 | section {* Intended Audience and Prior Knowledge *}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 21 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 22 | text {* 
 | 
| 75 | 23 | This Cookbook targets readers who already know how to use Isabelle for | 
| 24 | writing theories and proofs. We also assume that readers are familiar with | |
| 25 | the functional programming language ML, the language in which most of | |
| 26 | Isabelle is implemented. If you are unfamiliar with either of these two | |
| 27 | subjects, you should first work through the Isabelle/HOL tutorial | |
| 28 |   \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 | 29 | |
| 
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 | |
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 32 | section {* Existing Documentation *}
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 33 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 34 | text {*
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 35 | |
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 36 | 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: 
42diff
changeset | 37 | part of the distribution of Isabelle): | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 38 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 39 |   \begin{description}
 | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 40 | \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: 
43diff
changeset | 41 | from a high-level perspective, documenting both the underlying | 
| 6 | 42 | concepts and some of the interfaces. | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 43 | |
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 44 | \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: 
43diff
changeset | 45 | to be the main reference of Isabelle at a time when all proof scripts | 
| 89 | 46 | 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: 
43diff
changeset | 47 | 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: 
43diff
changeset | 48 | 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: 
43diff
changeset | 49 | |
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 50 | \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: 
43diff
changeset | 51 | 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: 
43diff
changeset | 52 | is still useful. | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 53 |   \end{description}
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 54 | |
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 55 | Then of course there is: | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 56 | |
| 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 57 |   \begin{description}
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 58 | \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 | 59 | 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 | 60 | way things are actually implemented. More importantly, it is often | 
| 89 | 61 | good to look at code that does similar things as you want to do and | 
| 62 | to learn from other people's code. | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 63 |   \end{description}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 64 | |
| 
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 | |
| 81 | 67 | section {* Typographic Conventions *}
 | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
66diff
changeset | 68 | |
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
66diff
changeset | 69 | text {*
 | 
| 75 | 70 | |
| 89 | 71 | All ML-code in this Cookbook is typeset in highlighted boxes, like the following | 
| 72 | ML-expression: | |
| 75 | 73 | |
| 74 |   \begin{isabelle}
 | |
| 75 |   \begin{graybox}
 | |
| 85 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
84diff
changeset | 76 |   \isacommand{ML}~@{text "\<verbopen>"}\isanewline
 | 
| 75 | 77 |   \hspace{5mm}@{ML "3 + 4"}\isanewline
 | 
| 85 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
84diff
changeset | 78 |   @{text "\<verbclose>"}
 | 
| 75 | 79 |   \end{graybox}
 | 
| 80 |   \end{isabelle}
 | |
| 81 | ||
| 89 | 82 | These boxes corresponds to how code can be processed inside the interactive | 
| 83 | environment of Isabelle. It is therefore easy to experiment with what is | |
| 84 | displayed. However, for better readability we will drop the enclosing | |
| 85 |   \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"} and just write:
 | |
| 75 | 86 | |
| 81 | 87 | |
| 75 | 88 |   @{ML [display,gray] "3 + 4"}
 | 
| 89 | ||
| 89 | 90 | Whenever appropriate we also show the response the code | 
| 81 | 91 | generates when evaluated. This response is prefixed with a | 
| 89 | 92 |   @{text [quotes] ">"} like:
 | 
| 75 | 93 | |
| 94 |   @{ML_response [display,gray] "3 + 4" "7"}
 | |
| 95 | ||
| 81 | 96 |   The usual Isabelle commands are written in bold, for example \isacommand{lemma}, 
 | 
| 80 | 97 |   \isacommand{foobar} and so on.  We use @{text "$"} to indicate that 
 | 
| 89 | 98 | a command needs to be run in the Unix-shell, for example | 
| 75 | 99 | |
| 100 |   @{text [display] "$ ls -la"}
 | |
| 101 | ||
| 89 | 102 | Pointers to further information and Isabelle files are typeset in | 
| 103 | italic and highlighted as follows: | |
| 75 | 104 | |
| 105 |   \begin{readmore}
 | |
| 80 | 106 | Further information or pointer to file. | 
| 75 | 107 |   \end{readmore}
 | 
| 108 | ||
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
66diff
changeset | 109 | *} | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
66diff
changeset | 110 | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 111 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 112 | end |