ProgTutorial/Intro.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 05 Aug 2009 16:00:01 +0200
changeset 306 fe732e890d87
parent 303 05e6a33edef6
child 324 4172c0743cf2
permissions -rw-r--r--
tuned the section about printing several bits of inormation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Intro
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
     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
303
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
     7
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
text {*
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
     9
   \begin{flushright}
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
    10
  {\em
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
    11
  ``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: 293
diff changeset
    12
  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: 293
diff changeset
    13
  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: 293
diff changeset
    14
  claimed, it's amongst the hardest intellectual tasks ever attempted.''} \\[1ex]
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
    15
  Richard Bornat, In Defence of Programming \cite{Bornat-lecture}
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
    16
  \end{flushright}
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
    17
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
    18
  \medskip
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    19
  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
    20
  then this tutorial is for you. It will guide you through the first steps of
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    21
  Isabelle programming, and also explain tricks of the trade. The best way to
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    22
  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
    23
  examples included in the tutorial. The code is as far as possible checked
254
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    24
  against the Isabelle distribution.\footnote{\input{version}} If something does not work, 
235
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    25
  then please let us know. It is impossible for us to know every environment, 
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    26
  operating system or editor in which Isabelle is used. If you have comments, 
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    27
  criticism or like to add to the tutorial, please feel free---you are most 
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    28
  welcome! The tutorial is meant to be gentle and comprehensive. To achieve 
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    29
  this we need your feedback.
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
section {* Intended Audience and Prior Knowledge *}
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 {* 
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
    35
  This tutorial targets readers who already know how to use Isabelle for
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    36
  writing theories and proofs. We also assume that readers are familiar with
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    37
  the functional programming language ML, the language in which most of
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    38
  Isabelle is implemented. If you are unfamiliar with either of these two
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    39
  subjects, you should first work through the Isabelle/HOL tutorial
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    40
  \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
    41
*}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    43
section {* Existing Documentation *}
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
text {*
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
    46
  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
    47
  part of the distribution of Isabelle):
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
  \begin{description}
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
    50
  \item[The Isabelle/Isar 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
    51
  from a high-level perspective, documenting both the underlying
6
007e09485351 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 5
diff changeset
    52
  concepts and some of the interfaces. 
2
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: 2
diff changeset
    54
  \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
    55
  to be the main reference of Isabelle at a time when all proof scripts 
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    56
  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
    57
  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
    58
  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
    59
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
    60
  \item[The Isar Reference Manual] provides specification material (like grammars,
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 295
diff changeset
    61
  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: 2
diff changeset
    62
  \end{description}
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
234
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
    64
  Then of course there are:
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    65
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    66
  \begin{description}
234
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
    67
  \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
    68
  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
    69
  way things are actually implemented. More importantly, it is often
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    70
  good to look at code that does similar things as you want to do and
254
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    71
  learn from it. The UNIX command \mbox{@{text "grep -R"}} is
233
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
    72
  often your best friend while programming with Isabelle, or
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
    73
  hypersearch if you program using jEdit under MacOSX. To understand the sources,
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
    74
  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: 262
diff changeset
    75
  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: 262
diff changeset
    76
  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: 262
diff changeset
    77
  files and ``change sets''.
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  \end{description}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
*}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
252
f380b13b25a7 added an example
Christian Urban <urbanc@in.tum.de>
parents: 248
diff changeset
    81
section {* Typographic Conventions *}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
    82
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
    83
text {*
181
5baaabe1ab92 updated to new method_setup
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
    84
  All ML-code in this tutorial is typeset in shaded boxes, like the following 
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    85
  ML-expression:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    86
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    87
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    88
  \begin{graybox}
85
b02904872d6b better handling of {* and *}
Christian Urban <urbanc@in.tum.de>
parents: 84
diff changeset
    89
  \isacommand{ML}~@{text "\<verbopen>"}\isanewline
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    90
  \hspace{5mm}@{ML "3 + 4"}\isanewline
85
b02904872d6b better handling of {* and *}
Christian Urban <urbanc@in.tum.de>
parents: 84
diff changeset
    91
  @{text "\<verbclose>"}
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    92
  \end{graybox}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    93
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    94
  
195
7305beb69893 corrected typos + some small reformulations
griff@colo2-c703.uibk.ac.at
parents: 192
diff changeset
    95
  These boxes correspond to how code can be processed inside the interactive
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    96
  environment of Isabelle. It is therefore easy to experiment with what is 
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    97
  displayed. However, for better readability we will drop the enclosing 
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    98
  \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"} and just write:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    99
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
   100
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   101
  @{ML [display,gray] "3 + 4"}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   102
  
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   103
  Whenever appropriate we also show the response the code 
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
   104
  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
   105
  @{text [quotes] ">"}, like:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   106
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   107
  @{ML_response [display,gray] "3 + 4" "7"}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   108
195
7305beb69893 corrected typos + some small reformulations
griff@colo2-c703.uibk.ac.at
parents: 192
diff changeset
   109
  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: 192
diff changeset
   110
  in \isacommand{bold face} (e.g., \isacommand{lemma}, \isacommand{apply},
7305beb69893 corrected typos + some small reformulations
griff@colo2-c703.uibk.ac.at
parents: 192
diff changeset
   111
  \isacommand{foobar} and so on).  We use @{text "$ \<dots>"} to indicate that a
234
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   112
  command needs to be run in a UNIX-shell, for example:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   113
181
5baaabe1ab92 updated to new method_setup
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   114
  @{text [display] "$ grep -R ThyOutput *"}
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   115
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   116
  Pointers to further information and Isabelle files are typeset in 
195
7305beb69893 corrected typos + some small reformulations
griff@colo2-c703.uibk.ac.at
parents: 192
diff changeset
   117
  \textit{italic} and highlighted as follows:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   118
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   119
  \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
   120
  Further information or pointers to files.
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   121
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   122
182
4d0e2edd476d added hyperlinks for every file pointer
Christian Urban <urbanc@in.tum.de>
parents: 181
diff changeset
   123
  The pointers to Isabelle files are hyperlinked to the tip of the Mercurial
4d0e2edd476d added hyperlinks for every file pointer
Christian Urban <urbanc@in.tum.de>
parents: 181
diff changeset
   124
  repository of Isabelle at \href{http://isabelle.in.tum.de/repos/isabelle/}
4d0e2edd476d added hyperlinks for every file pointer
Christian Urban <urbanc@in.tum.de>
parents: 181
diff changeset
   125
  {http://isabelle.in.tum.de/repos/isabelle/}.
4d0e2edd476d added hyperlinks for every file pointer
Christian Urban <urbanc@in.tum.de>
parents: 181
diff changeset
   126
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
   127
  A few exercises are scattered around the text. Their solutions are given 
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   128
  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: 156
diff changeset
   129
  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: 228
diff changeset
   130
*}
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   131
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   132
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: 262
diff changeset
   133
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   134
text {*
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   135
  One unpleasant aspect of any code development inside a larger system is that
264
311830b43f8f minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   136
  one has to aim at a ``moving target''. The Isabelle system is no exception. Every 
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   137
  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: 262
diff changeset
   138
  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: 262
diff changeset
   139
  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: 262
diff changeset
   140
  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: 262
diff changeset
   141
  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: 262
diff changeset
   142
  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: 262
diff changeset
   143
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   144
  However, there are a few steps you can take to mitigate unwanted interferences
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   145
  with code changes from other developers. First, you can base your code on the latest 
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   146
  stable release of Isabelle (it is aimed to have one such release at least 
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   147
  once every year). This 
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   148
  might cut you off from the latest feature
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   149
  implemented in Isabelle, but at least you do not have to track side-steps
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   150
  or dead-ends in the Isabelle development. Of course this means also you have
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   151
  to synchronise your code at the next stable release. Another possibility
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   152
  is to get your code into the Isabelle distribution. For this you have
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   153
  to convince other developers that your code or project is of general interest. 
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   154
  If you managed to do this, then the problem of the moving target goes 
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 295
diff changeset
   155
  away, because when checking in new code, developers are strongly urged to 
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 295
diff changeset
   156
  test it against Isabelle's code base. If your project is part of that code base, 
265
c354e45d80d2 small changes for latest changes in Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 264
diff changeset
   157
  then maintenance is done by others. Unfortunately, this might not
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   158
  be a helpful advice for all types of projects. A lower threshold for inclusion has the 
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   159
  Archive of Formalised Proofs, short AFP.\footnote{\url{http://afp.sourceforge.net/}}
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   160
  This archive has been created mainly for formalisations that are
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   161
  interesting but not necessarily of general interest. If you have ML-code as
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   162
  part of a formalisation, then this might be the right place for you. There
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   163
  is no problem with updating your code after submission. At the moment
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   164
  developers are not as diligent with checking their code against the AFP than
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   165
  with checking agains the distribution, but generally problems will be caught
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   166
  and the developer, who caused them, is expected to fix them. So also
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   167
  in this case code maintenance is done for you. 
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 295
diff changeset
   168
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   169
*}
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   170
252
f380b13b25a7 added an example
Christian Urban <urbanc@in.tum.de>
parents: 248
diff changeset
   171
section {* Some Naming Conventions in the Isabelle Sources *}
233
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   172
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   173
text {*
254
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
   174
  There are a few naming conventions in the Isabelle code that might aid reading 
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
   175
  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: 228
diff changeset
   176
  times.) The most important conventions are:
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   177
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   178
  \begin{itemize}
302
Christian Urban <urbanc@in.tum.de>
parents: 298
diff changeset
   179
  \item @{text t}, @{text u}, @{text trm} for (raw) terms; ML-type: @{ML_type term}
234
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   180
  \item @{text ct}, @{text cu} for certified terms; ML-type: @{ML_type cterm}
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   181
  \item @{text "ty"}, @{text T}, @{text U} for (raw) types; ML-type: @{ML_type typ}
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   182
  \item @{text th}, @{text thm} for theorems; ML-type: @{ML_type thm}
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   183
  \item @{text "foo_tac"} for tactics; ML-type: @{ML_type tactic}
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   184
  \item @{text thy} for theories; ML-type: @{ML_type theory}
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   185
  \item @{text ctxt} for proof contexts; ML-type: @{ML_type Proof.context}
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   186
  \item @{text lthy} for local theories; ML-type: @{ML_type local_theory}
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   187
  \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: 234
diff changeset
   188
  \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: 235
diff changeset
   189
  \item @{text prt} for pretty printing; ML-type @{ML_type Pretty.T}
233
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   190
  \end{itemize}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
   191
*}
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
   192
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   193
section {* Acknowledgements *}
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   194
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   195
text {*
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   196
  Financial support for this tutorial was provided by the German 
122
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   197
  Research Council (DFG) under grant number URB 165/5-1. The following
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   198
  people contributed to the text:
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   199
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   200
  \begin{itemize}
122
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   201
  \item {\bf Stefan Berghofer} wrote nearly all of the ML-code of the
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   202
  \simpleinductive-package and the code for the @{text
264
311830b43f8f minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   203
  "chunk"}-antiquotation. He also wrote the first version of this chapter
122
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   204
  describing the package and has been helpful \emph{beyond measure} with
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   205
  answering questions about Isabelle.
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   206
248
11851b20fb78 added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
parents: 235
diff changeset
   207
  \item {\bf Jasmin Blanchette} helped greatly with section \ref{sec:pretty}.
11851b20fb78 added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
parents: 235
diff changeset
   208
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   209
  \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
   210
  \ref{rec:config}, \ref{rec:storingdata}, \ref{rec:external} and \ref{rec:oracle}.
167
3e30ea95c7aa added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
   211
  He also wrote section \ref{sec:conversion} and helped with recipe \ref{rec:timing}.
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   212
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   213
  \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
   214
  about parsing.
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   215
180
9c25418db6f0 added a recipy about SAT solvers
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   216
  \item {\bf Armin Heller} helped with recipe \ref{rec:sat}.
9c25418db6f0 added a recipy about SAT solvers
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   217
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   218
  \item {\bf Alexander Krauss} wrote the first version of the ``first-steps'' 
293
0a567f923b42 slightly changed exercises about rev_sum
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   219
  chapter and also contributed the material on @{ML_functor Named_Thms}.
194
8cd51a25a7ca some polishing
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   220
207
d3cd633e8240 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 203
diff changeset
   221
  \item {\bf Christian Sternagel} proofread the tutorial and made 
293
0a567f923b42 slightly changed exercises about rev_sum
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   222
  many improvemets to the text. 
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   223
  \end{itemize}
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   224
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   225
  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: 151
diff changeset
   226
  errors lies with me.\bigskip
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   227
233
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   228
  \vspace{5cm}
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   229
  {\Large\bf
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   230
  This document is still in the process of being written! All of the
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   231
  text is still under construction. Sections and 
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   232
  chapters that are under \underline{heavy} construction are marked 
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   233
  with TBD.}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   234
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   235
  \vfill
306
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 303
diff changeset
   236
  
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 303
diff changeset
   237
  This document (version \input{tip}\hspace{-0.5ex}) was compiled with:\\
228
fe45fbb111c5 various additions
Christian Urban <urbanc@in.tum.de>
parents: 207
diff changeset
   238
  \input{version}\\
fe45fbb111c5 various additions
Christian Urban <urbanc@in.tum.de>
parents: 207
diff changeset
   239
  \input{pversion}
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   240
*}
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   242
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   243
195
7305beb69893 corrected typos + some small reformulations
griff@colo2-c703.uibk.ac.at
parents: 192
diff changeset
   244
end