| author | Christian Urban <christian dot urban at kcl dot ac dot uk> | 
| Mon, 25 Feb 2013 00:33:48 +0000 | |
| changeset 542 | 4b96e3c8b33e | 
| parent 540 | d144fc51fe04 | 
| child 553 | c53d74b34123 | 
| 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 | begin | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 5 | chapter {* Introduction *}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 6 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 7 | text {*
 | 
| 295 
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
 Christian Urban <urbanc@in.tum.de> parents: 
293diff
changeset | 8 |    \begin{flushright}
 | 
| 
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
 Christian Urban <urbanc@in.tum.de> parents: 
293diff
changeset | 9 |   {\em
 | 
| 
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
 Christian Urban <urbanc@in.tum.de> parents: 
293diff
changeset | 10 | ``My thesis is that programming is not at the bottom of the intellectual \\ | 
| 
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
 Christian Urban <urbanc@in.tum.de> parents: 
293diff
changeset | 11 | pyramid, but at the top. It's creative design of the highest order. It \\ | 
| 
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
 Christian Urban <urbanc@in.tum.de> parents: 
293diff
changeset | 12 | isn't monkey or donkey work; rather, as Edsger Dijkstra famously \\ | 
| 
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
 Christian Urban <urbanc@in.tum.de> parents: 
293diff
changeset | 13 | claimed, it's amongst the hardest intellectual tasks ever attempted.''} \\[1ex] | 
| 324 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 14 |   Richard Bornat, In {\em Defence of Programming}. \cite{Bornat-lecture}
 | 
| 295 
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
 Christian Urban <urbanc@in.tum.de> parents: 
293diff
changeset | 15 |   \end{flushright}
 | 
| 
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
 Christian Urban <urbanc@in.tum.de> parents: 
293diff
changeset | 16 | |
| 
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
 Christian Urban <urbanc@in.tum.de> parents: 
293diff
changeset | 17 | \medskip | 
| 89 | 18 | 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: 
102diff
changeset | 19 | then this tutorial is for you. It will guide you through the first steps of | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
435diff
changeset | 20 | Isabelle programming, and also explain ``tricks of the trade''. We also hope | 
| 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
435diff
changeset | 21 | the tutorial will encourage students and researchers to play with Isabelle | 
| 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
435diff
changeset | 22 | and implement new ideas. The source code of Isabelle can look intimidating, | 
| 466 | 23 | but beginners can get by with knowledge of only a handful of concepts, | 
| 24 | a small number of functions and a few basic coding conventions. | |
| 356 | 25 | |
| 26 | ||
| 27 | The best way to get to know the ML-level of Isabelle is by experimenting | |
| 28 | with the many code examples included in the tutorial. The code is as far as | |
| 29 | possible checked against the Isabelle | |
| 540 
d144fc51fe04
removed versions
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
536diff
changeset | 30 |   distribution.%%\footnote{\input{version.tex}}
 | 
| 535 
5734ab5dd86d
adapted to new build framework
 Christian Urban <urbanc@in.tum.de> parents: 
532diff
changeset | 31 | If something does not work, then | 
| 356 | 32 | please let us know. It is impossible for us to know every environment, | 
| 33 | operating system or editor in which Isabelle is used. If you have comments, | |
| 34 | criticism or like to add to the tutorial, please feel free---you are most | |
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
435diff
changeset | 35 | welcome!! The tutorial is meant to be gentle and comprehensive. To achieve | 
| 414 | 36 | this we need your help and feedback. | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 37 | *} | 
| 
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 | section {* Intended Audience and Prior Knowledge *}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 40 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 41 | 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: 
102diff
changeset | 42 | This tutorial targets readers who already know how to use Isabelle for | 
| 75 | 43 | writing theories and proofs. We also assume that readers are familiar with | 
| 44 | the functional programming language ML, the language in which most of | |
| 45 | Isabelle is implemented. If you are unfamiliar with either of these two | |
| 329 | 46 | subjects, then you should first work through the Isabelle/HOL tutorial | 
| 466 | 47 |   \cite{isa-tutorial} or Paulson's book on ML \cite{paulson-ml2}. Recently,
 | 
| 48 | Isabelle has adopted a sizable amount of Scala code for a slick GUI | |
| 49 | based on jEdit. This part of the code is beyond the interest of this | |
| 50 | tutorial, since it mostly does not concern the regular Isabelle | |
| 51 | developer. | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 52 | *} | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 53 | |
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 54 | section {* Existing Documentation *}
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 55 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 56 | text {*
 | 
| 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 | 57 | 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 | 58 | part of the distribution of Isabelle): | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 59 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 60 |   \begin{description}
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 61 | \item[The Isabelle/Isar Implementation Manual] describes Isabelle | 
| 414 | 62 | from a high-level perspective, documenting some of the underlying | 
| 63 | concepts and interfaces. | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 64 | |
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 65 | \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 | 66 | to be the main reference of Isabelle at a time when all proof scripts | 
| 89 | 67 | 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 | 68 | 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 | 69 | 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 | 70 | |
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 71 | \item[The Isar Reference Manual] provides specification material (like grammars, | 
| 298 | 72 | examples and so on) about Isar and its implementation. | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 73 |   \end{description}
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 74 | |
| 234 | 75 | Then of course there are: | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 76 | |
| 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 77 |   \begin{description}
 | 
| 234 | 78 | \item[The Isabelle sources.] They are the ultimate reference for how | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 79 | things really work. Therefore you should not hesitate to look at the | 
| 466 | 80 | way things are actually implemented. While much of the Isabelle | 
| 81 | code is uncommented, some parts have very helpful comments---particularly | |
| 82 | the code about theorems and terms. Despite the lack of comments in most | |
| 83 | parts, it is often good to look at code that does similar things as you | |
| 84 | want to do and learn from it. | |
| 85 | This tutorial contains frequently pointers to the | |
| 329 | 86 |   Isabelle sources. Still, the UNIX command \mbox{@{text "grep -R"}} is
 | 
| 414 | 87 |   often your best friend while programming with Isabelle.\footnote{Or
 | 
| 441 | 88 | hypersearch if you work with jEdit.} To understand the sources, | 
| 263 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 89 | it is often also necessary to track the change history of a file or | 
| 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 90 |   files. The Mercurial repository\footnote{\url{http://isabelle.in.tum.de/repos/isabelle/}} 
 | 
| 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 91 | for Isabelle provides convenient interfaces to query the history of | 
| 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 92 | files and ``change sets''. | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 93 |   \end{description}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 94 | *} | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 95 | |
| 252 | 96 | section {* Typographic Conventions *}
 | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
66diff
changeset | 97 | |
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
66diff
changeset | 98 | text {*
 | 
| 181 
5baaabe1ab92
updated to new method_setup
 Christian Urban <urbanc@in.tum.de> parents: 
180diff
changeset | 99 | All ML-code in this tutorial is typeset in shaded boxes, like the following | 
| 329 | 100 | simple ML-expression: | 
| 75 | 101 | |
| 102 |   \begin{isabelle}
 | |
| 103 |   \begin{graybox}
 | |
| 85 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
84diff
changeset | 104 |   \isacommand{ML}~@{text "\<verbopen>"}\isanewline
 | 
| 75 | 105 |   \hspace{5mm}@{ML "3 + 4"}\isanewline
 | 
| 85 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
84diff
changeset | 106 |   @{text "\<verbclose>"}
 | 
| 75 | 107 |   \end{graybox}
 | 
| 108 |   \end{isabelle}
 | |
| 109 | ||
| 195 
7305beb69893
corrected typos + some small reformulations
 griff@colo2-c703.uibk.ac.at parents: 
192diff
changeset | 110 | These boxes correspond to how code can be processed inside the interactive | 
| 343 | 111 | environment of Isabelle. It is therefore easy to experiment with the code | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
435diff
changeset | 112 | that is shown in this tutorial. However, for better readability we will drop | 
| 343 | 113 |   the enclosing \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"} and just
 | 
| 114 | write: | |
| 81 | 115 | |
| 75 | 116 |   @{ML [display,gray] "3 + 4"}
 | 
| 117 | ||
| 89 | 118 | Whenever appropriate we also show the response the code | 
| 81 | 119 | 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: 
107diff
changeset | 120 |   @{text [quotes] ">"}, like:
 | 
| 75 | 121 | |
| 122 |   @{ML_response [display,gray] "3 + 4" "7"}
 | |
| 123 | ||
| 195 
7305beb69893
corrected typos + some small reformulations
 griff@colo2-c703.uibk.ac.at parents: 
192diff
changeset | 124 | The user-level commands of Isabelle (i.e., the non-ML code) are written | 
| 
7305beb69893
corrected typos + some small reformulations
 griff@colo2-c703.uibk.ac.at parents: 
192diff
changeset | 125 |   in \isacommand{bold face} (e.g., \isacommand{lemma}, \isacommand{apply},
 | 
| 
7305beb69893
corrected typos + some small reformulations
 griff@colo2-c703.uibk.ac.at parents: 
192diff
changeset | 126 |   \isacommand{foobar} and so on).  We use @{text "$ \<dots>"} to indicate that a
 | 
| 234 | 127 | command needs to be run in a UNIX-shell, for example: | 
| 75 | 128 | |
| 441 | 129 |   @{text [display] "$ grep -R Thy_Output *"}
 | 
| 75 | 130 | |
| 89 | 131 | Pointers to further information and Isabelle files are typeset in | 
| 195 
7305beb69893
corrected typos + some small reformulations
 griff@colo2-c703.uibk.ac.at parents: 
192diff
changeset | 132 |   \textit{italic} and highlighted as follows:
 | 
| 75 | 133 | |
| 134 |   \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: 
101diff
changeset | 135 | Further information or pointers to files. | 
| 75 | 136 |   \end{readmore}
 | 
| 137 | ||
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
435diff
changeset | 138 | Note that pointers to Isabelle files are hyperlinked to the tip of the Mercurial | 
| 414 | 139 |   repository at \href{http://isabelle.in.tum.de/repos/isabelle/}
 | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
435diff
changeset | 140 |   {http://isabelle.in.tum.de/repos/isabelle/}, not the latest stable release
 | 
| 414 | 141 | of Isabelle. | 
| 182 
4d0e2edd476d
added hyperlinks for every file pointer
 Christian Urban <urbanc@in.tum.de> parents: 
181diff
changeset | 142 | |
| 177 | 143 | A few exercises are scattered around the text. Their solutions are given | 
| 156 | 144 |   in Appendix~\ref{ch:solutions}. Of course, you learn most, if you first try
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 145 | to solve the exercises on your own, and then look at the solutions. | 
| 233 
61085dd44e8c
added a section about naming conventions
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 146 | *} | 
| 156 | 147 | |
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
435diff
changeset | 148 | section {* How To Understand Isabelle Code *}
 | 
| 421 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
417diff
changeset | 149 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
417diff
changeset | 150 | text {*
 | 
| 466 | 151 | One of the more difficult aspects of any kind of programming is to | 
| 152 | understand code written by somebody else. This is aggravated in Isabelle by | |
| 153 | the fact that many parts of the code contain none or only few | |
| 154 | comments. There is one strategy that might be helpful to navigate your way: | |
| 155 | ML is an interactive programming environment, which means you can evaluate | |
| 156 |   code on the fly (for example inside an \isacommand{ML}~@{text
 | |
| 157 | "\<verbopen>\<dots>\<verbclose>"} section). So you can copy (self-contained) | |
| 158 | chunks of existing code into a separate theory file and then study it | |
| 159 | alongside with examples. You can also install ``probes'' inside the copied | |
| 160 | code without having to recompile the whole Isabelle distribution. Such | |
| 421 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
417diff
changeset | 161 | probes might be messages or printouts of variables (see chapter | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
417diff
changeset | 162 |   \ref{chp:firststeps}). Although PolyML also contains a debugger, it seems
 | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
435diff
changeset | 163 | probing the code with explicit print statements is the most effective method | 
| 466 | 164 | for understanding what some piece of code is doing. However do not expect | 
| 165 | quick results with this! It is painful. Depending on the size of the code | |
| 166 | you are looking at, you will spend the better part of a quiet afternoon with | |
| 167 | it. And there seems to be no better way for understanding code in Isabelle. | |
| 421 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
417diff
changeset | 168 | *} | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
417diff
changeset | 169 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
417diff
changeset | 170 | |
| 263 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 171 | section {* Aaaaargh! My Code Does not Work Anymore *}
 | 
| 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 172 | |
| 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 173 | text {*
 | 
| 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 174 | One unpleasant aspect of any code development inside a larger system is that | 
| 440 
a0b280dd4bc7
partially moved from string_of_term to pretty_term
 Christian Urban <urbanc@in.tum.de> parents: 
435diff
changeset | 175 | one has to aim at a ``moving target''. Isabelle is no exception of this. Every | 
| 263 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 176 | update lets potentially all hell break loose, because other developers have | 
| 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 177 | changed code you are relying on. Cursing is somewhat helpful in such situations, | 
| 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 178 | but taking the view that incompatible code changes are a fact of life | 
| 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 179 | might be more gratifying. Isabelle is a research project. In most circumstances | 
| 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 180 | it is just impossible to make research backward compatible (imagine Darwin | 
| 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 181 | attempting to make the Theory of Evolution backward compatible). | 
| 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 182 | |
| 329 | 183 | However, there are a few steps you can take to mitigate unwanted | 
| 184 | interferences with code changes from other developers. First, you can base | |
| 185 | your code on the latest stable release of Isabelle (it is aimed to have one | |
| 186 | such release at least once every year). This might cut you off from the | |
| 187 | latest feature implemented in Isabelle, but at least you do not have to | |
| 188 | track side-steps or dead-ends in the Isabelle development. Of course this | |
| 189 | means also you have to synchronise your code at the next stable release. If | |
| 190 | you do not synchronise, be warned that code seems to ``rot'' very | |
| 191 | quickly. Another possibility is to get your code into the Isabelle | |
| 192 | distribution. For this you have to convince other developers that your code | |
| 193 | or project is of general interest. If you managed to do this, then the | |
| 194 | problem of the moving target goes away, because when checking in new code, | |
| 195 | developers are strongly urged to test it against Isabelle's code base. If | |
| 196 | your project is part of that code base, then maintenance is done by | |
| 197 | others. Unfortunately, this might not be a helpful advice for all types of | |
| 454 | 198 | projects. A lower threshold for inclusion has the Archive of Formal | 
| 329 | 199 |   Proofs, short AFP.\footnote{\url{http://afp.sourceforge.net/}} This archive
 | 
| 200 | has been created mainly for formalisations that are interesting but not | |
| 201 | necessarily of general interest. If you have ML-code as part of a | |
| 202 | formalisation, then this might be the right place for you. There is no | |
| 203 | problem with updating your code after submission. At the moment developers | |
| 204 | are not as diligent with checking their code against the AFP than with | |
| 205 | checking agains the distribution, but generally problems will be caught and | |
| 206 | the developer, who caused them, is expected to fix them. So also in this | |
| 207 | case code maintenance is done for you. | |
| 263 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 208 | *} | 
| 
195c4444dff7
added section about code maintenance and added an example for antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
262diff
changeset | 209 | |
| 485 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 210 | section {* Serious Isabelle ML-Programming *}
 | 
| 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 211 | |
| 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 212 | text {*
 | 
| 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 213 | As already pointed out in the previous section, Isabelle is a joint effort | 
| 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 214 | of many developers. Therefore, disruptions that break the work of others | 
| 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 215 | are generally frowned upon. ``Accidents'' however do happen and everybody knows | 
| 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 216 | this. Still to keep them to a minimum, you can submit your changes first to a rather | 
| 486 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 217 |   sophisticated \emph{testboard}, which will perform checks of your changes against the
 | 
| 485 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 218 | Isabelle repository and against the AFP. The advantage of the testboard is | 
| 502 | 219 | that the testing is performed by rather powerful machines saving you lengthy | 
| 486 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 220 | tests on, for example, your own laptop. You can see the results of the testboard | 
| 485 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 221 | at | 
| 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 222 | |
| 486 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 223 |   \begin{center}
 | 
| 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 224 |   \url{http://isabelle.in.tum.de/testboard/Isabelle/}
 | 
| 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 225 |   \end{center}
 | 
| 485 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 226 | |
| 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 227 | which is organised like a Mercurial repository. A green point next to a change | 
| 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 228 | indicates that the change passes the corresponding tests (for this of course you | 
| 486 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 229 | have to allow some time). You can summit any changes to the testboard using the | 
| 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 230 | command | 
| 485 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 231 | |
| 532 | 232 |   @{text [display] "$ hg push -f ssh://...@hgbroy.informatik.tu-muenchen.de\\
 | 
| 486 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 233 | //home/isabelle-repository/repos/testboard"} | 
| 485 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 234 | |
| 486 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 235 | where the dots need to be replaced by your login name. Note that for | 
| 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 236 |   pushing changes to the testboard you need to add the option @{text "-f"},
 | 
| 502 | 237 |   which should \emph{never} be used with the main Isabelle
 | 
| 486 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 238 | repository. While the testboard is a great system for supporting Isabelle | 
| 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 239 | developers, its disadvantage is that it needs login permissions for the | 
| 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 240 | computers in Munich. So in order to use it, you might have to ask other | 
| 
45cfd2ece7bd
a section about theories and setups
 Christian Urban <urbanc@in.tum.de> parents: 
485diff
changeset | 241 | developers to obtain one. | 
| 485 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 242 | *} | 
| 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 243 | |
| 
f3536f0b47a9
added section about testboard
 Christian Urban <urbanc@in.tum.de> parents: 
469diff
changeset | 244 | |
| 252 | 245 | section {* Some Naming Conventions in the Isabelle Sources *}
 | 
| 233 
61085dd44e8c
added a section about naming conventions
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 246 | |
| 
61085dd44e8c
added a section about naming conventions
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 247 | text {*
 | 
| 254 | 248 | There are a few naming conventions in the Isabelle code that might aid reading | 
| 249 | and writing code. (Remember that code is written once, but read many | |
| 233 
61085dd44e8c
added a section about naming conventions
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 250 | times.) The most important conventions are: | 
| 
61085dd44e8c
added a section about naming conventions
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 251 | |
| 
61085dd44e8c
added a section about naming conventions
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 252 |   \begin{itemize}
 | 
| 302 | 253 |   \item @{text t}, @{text u}, @{text trm} for (raw) terms; ML-type: @{ML_type term}
 | 
| 234 | 254 |   \item @{text ct}, @{text cu} for certified terms; ML-type: @{ML_type cterm}
 | 
| 255 |   \item @{text "ty"}, @{text T}, @{text U} for (raw) types; ML-type: @{ML_type typ}
 | |
| 390 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
358diff
changeset | 256 |   \item @{text "S"} for sorts; ML-type: @{ML_type sort}
 | 
| 234 | 257 |   \item @{text th}, @{text thm} for theorems; ML-type: @{ML_type thm}
 | 
| 258 |   \item @{text "foo_tac"} for tactics; ML-type: @{ML_type tactic}
 | |
| 259 |   \item @{text thy} for theories; ML-type: @{ML_type theory}
 | |
| 260 |   \item @{text ctxt} for proof contexts; ML-type: @{ML_type Proof.context}
 | |
| 261 |   \item @{text lthy} for local theories; ML-type: @{ML_type local_theory}
 | |
| 262 |   \item @{text context} for generic contexts; ML-type @{ML_type Context.generic}
 | |
| 235 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 263 |   \item @{text mx} for mixfix syntax annotations; ML-type @{ML_type mixfix}
 | 
| 248 
11851b20fb78
added more to the pretty section and updated the acknowledgements
 Christian Urban <urbanc@in.tum.de> parents: 
235diff
changeset | 264 |   \item @{text prt} for pretty printing; ML-type @{ML_type Pretty.T}
 | 
| 425 | 265 |   \item @{text phi} for morphisms; ML-type @{ML_type morphism}
 | 
| 233 
61085dd44e8c
added a section about naming conventions
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 266 |   \end{itemize}
 | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
66diff
changeset | 267 | *} | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
66diff
changeset | 268 | |
| 119 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 269 | section {* Acknowledgements *}
 | 
| 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 270 | |
| 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 271 | text {*
 | 
| 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 272 | Financial support for this tutorial was provided by the German | 
| 122 | 273 | Research Council (DFG) under grant number URB 165/5-1. The following | 
| 156 | 274 | people contributed to the text: | 
| 119 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 275 | |
| 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 276 |   \begin{itemize}
 | 
| 122 | 277 |   \item {\bf Stefan Berghofer} wrote nearly all of the ML-code of the
 | 
| 278 |   \simpleinductive-package and the code for the @{text
 | |
| 414 | 279 | "chunk"}-antiquotation. He also wrote the first version of chapter | 
| 280 |   \ref{chp:package} describing this package and has been helpful \emph{beyond
 | |
| 281 | measure} with answering questions about Isabelle. | |
| 119 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 282 | |
| 469 
7a558c5119b2
added an excercise originally by Jasmin Blanchette
 Christian Urban <urbanc@in.tum.de> parents: 
466diff
changeset | 283 |   \item {\bf Jasmin Blanchette} helped greatly with section \ref{sec:pretty}
 | 
| 
7a558c5119b2
added an excercise originally by Jasmin Blanchette
 Christian Urban <urbanc@in.tum.de> parents: 
466diff
changeset | 284 |   and exercise \ref{fun:killqnt}.
 | 
| 248 
11851b20fb78
added more to the pretty section and updated the acknowledgements
 Christian Urban <urbanc@in.tum.de> parents: 
235diff
changeset | 285 | |
| 119 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 286 |   \item {\bf Sascha Böhme} contributed the recipes in \ref{rec:timeout}, 
 | 
| 328 
c0cae24b9d46
updated to new Isabelle; more work on the data section
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 287 |   \ref{rec:external} and \ref{rec:oracle}. He also wrote section \ref{sec:conversion} 
 | 
| 329 | 288 |   and helped with recipe \ref{rec:timing}. Parts of section \ref{sec:storing}
 | 
| 328 
c0cae24b9d46
updated to new Isabelle; more work on the data section
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 289 | are by him. | 
| 119 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 290 | |
| 414 | 291 |   \item {\bf Lukas Bulwahn} made me aware of a problem with recursive
 | 
| 530 | 292 |   parsers, contributed exercise \ref{ex:contextfree} and contributed 
 | 
| 293 |   to the ``introspection'' of theorems in section \ref{sec:theorems}.
 | |
| 529 
13d7ea419c5f
moved the introspection part into the theorem section
 Christian Urban <urbanc@in.tum.de> parents: 
526diff
changeset | 294 | |
| 390 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
358diff
changeset | 295 | |
| 414 | 296 |   \item {\bf Jeremy Dawson} wrote the first version of chapter \ref{chp:parsing}
 | 
| 119 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 297 | about parsing. | 
| 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 298 | |
| 542 
4b96e3c8b33e
updated the CallML section with the help from Florian
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
540diff
changeset | 299 |   \item {\bf Florian Haftmann} helped with maintaining recipe \ref{rec:callml}.
 | 
| 
4b96e3c8b33e
updated the CallML section with the help from Florian
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
540diff
changeset | 300 | |
| 180 
9c25418db6f0
added a recipy about SAT solvers
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 301 |   \item {\bf Armin Heller} helped with recipe \ref{rec:sat}.
 | 
| 
9c25418db6f0
added a recipy about SAT solvers
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 302 | |
| 530 | 303 |   \item {\bf Rafal Kolanski} contributed to the ``introspection'' of theorems 
 | 
| 304 |   in section \ref{sec:theorems}.
 | |
| 526 | 305 | |
| 414 | 306 |   \item {\bf Alexander Krauss} wrote a very early version of the ``first-steps'' 
 | 
| 343 | 307 |   chapter and also contributed the material on @{ML_funct Named_Thms}.
 | 
| 194 | 308 | |
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: 
425diff
changeset | 309 |   \item {\bf Tobias Nipkow} contributed recipe \ref{rec:callml}.
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: 
425diff
changeset | 310 | |
| 417 
5f00958e3c7b
typos fixed by Michael Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 311 |   \item {\bf Michael Norrish} proofread parts of the text.
 | 
| 
5f00958e3c7b
typos fixed by Michael Norrish
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 312 | |
| 435 | 313 |   \item {\bf Andreas Schropp} improved and corrected section \ref{sec:univ} and
 | 
| 314 |    contributed towards section \ref{sec:sorts}.
 | |
| 315 | ||
| 207 | 316 |   \item {\bf Christian Sternagel} proofread the tutorial and made 
 | 
| 293 
0a567f923b42
slightly changed exercises about rev_sum
 Christian Urban <urbanc@in.tum.de> parents: 
265diff
changeset | 317 | many improvemets to the text. | 
| 522 | 318 | |
| 319 |   \item {\bf Dmitriy Traytel} suggested to use the ML-antiquotation 
 | |
| 320 |   @{text "command_spec"} in section~\ref{sec:newcommand}, which simplified the code. 
 | |
| 119 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 321 |   \end{itemize}
 | 
| 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 322 | |
| 121 | 323 | Please let me know of any omissions. Responsibility for any remaining | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 324 | errors lies with me.\bigskip | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 325 | |
| 421 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
417diff
changeset | 326 | \newpage | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
417diff
changeset | 327 |   \mbox{}\\[5cm]
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
417diff
changeset | 328 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
417diff
changeset | 329 | |
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 330 |   {\Large\bf
 | 
| 335 
163ac0662211
reorganised the certified terms section; tuned
 Christian Urban <urbanc@in.tum.de> parents: 
329diff
changeset | 331 | This tutorial is still in the process of being written! All of the | 
| 192 | 332 | text is still under construction. Sections and | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 333 |   chapters that are under \underline{heavy} construction are marked 
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 334 | with TBD.} | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 335 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 336 | \vfill | 
| 540 
d144fc51fe04
removed versions
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
536diff
changeset | 337 |   %%This document (version \input{tip.tex}\hspace{-0.5ex}) was compiled with:\\
 | 
| 
d144fc51fe04
removed versions
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
536diff
changeset | 338 |   %%\input{version.tex}\\
 | 
| 535 
5734ab5dd86d
adapted to new build framework
 Christian Urban <urbanc@in.tum.de> parents: 
532diff
changeset | 339 |   %% \input{pversion}
 | 
| 119 
4536782969fa
added an acknowledgement section
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 340 | *} | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 341 | |
| 195 
7305beb69893
corrected typos + some small reformulations
 griff@colo2-c703.uibk.ac.at parents: 
192diff
changeset | 342 | end |