ProgTutorial/Intro.thy
author griff
Tue, 21 Jul 2009 11:53:41 +0200
changeset 267 83abec907072
parent 235 dc955603d813
child 248 11851b20fb78
permissions -rw-r--r--
if n would be 0, rev_sum would be missing an argument
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
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
text {*
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
     8
  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
     9
  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
    10
  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
    11
  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
    12
  examples included in the tutorial. The code is as far as possible checked
235
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    13
  against \input{version}. If something does not work, 
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    14
  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
    15
  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
    16
  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
    17
  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
    18
  this we need your feedback.
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
*}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
section {* Intended Audience and Prior Knowledge *}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
text {* 
106
bdd82350cf22 renamed in the pdf all instances of cookbook to tutorial (in order to sound more serious)
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    24
  This tutorial targets readers who already know how to use Isabelle for
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    25
  writing theories and proofs. We also assume that readers are familiar with
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    26
  the functional programming language ML, the language in which most of
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    27
  Isabelle is implemented. If you are unfamiliar with either of these two
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    28
  subjects, you should first work through the Isabelle/HOL tutorial
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    29
  \cite{isa-tutorial} or Paulson's book on ML \cite{paulson-ml2}.
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
*}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    33
section {* Existing Documentation *}
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
text {*
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  
43
02f76f1b6e7b added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
    37
  The following documentation about Isabelle programming already exists (and is
02f76f1b6e7b added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
    38
  part of the distribution of Isabelle):
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  \begin{description}
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
    41
  \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
    42
  from a high-level perspective, documenting both the underlying
6
007e09485351 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 5
diff changeset
    43
  concepts and some of the interfaces. 
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    45
  \item[The Isabelle Reference Manual] is an older document that used
44
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
    46
  to be the main reference of Isabelle at a time when all proof scripts 
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    47
  were written on the ML-level. Many parts of this manual are outdated 
44
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
    48
  now, but some  parts, particularly the chapters on tactics, are still 
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
    49
  useful.
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
    50
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
    51
  \item[The Isar Reference Manual] provides specification material (like grammars,
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
    52
  examples and so on) about Isar and its implementation. It is currently in
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
    53
  the process of being updated.
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    54
  \end{description}
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
234
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
    56
  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
    57
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    58
  \begin{description}
234
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
    59
  \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
    60
  things really work. Therefore you should not hesitate to look at the
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
  way things are actually implemented. More importantly, it is often
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    62
  good to look at code that does similar things as you want to do and
235
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    63
  learn from it. The GNU/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
    64
  often your best friend while programming with Isabelle, or
235
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    65
  hypersearch if you program using jEdit under MacOSX.
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
  \end{description}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
*}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
233
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
    70
section {* Typographic Conventions in the Tutorial *}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
    71
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
    72
text {*
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    73
181
5baaabe1ab92 updated to new method_setup
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
    74
  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
    75
  ML-expression:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    76
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    77
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    78
  \begin{graybox}
85
b02904872d6b better handling of {* and *}
Christian Urban <urbanc@in.tum.de>
parents: 84
diff changeset
    79
  \isacommand{ML}~@{text "\<verbopen>"}\isanewline
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    80
  \hspace{5mm}@{ML "3 + 4"}\isanewline
85
b02904872d6b better handling of {* and *}
Christian Urban <urbanc@in.tum.de>
parents: 84
diff changeset
    81
  @{text "\<verbclose>"}
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    82
  \end{graybox}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    83
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    84
  
195
7305beb69893 corrected typos + some small reformulations
griff@colo2-c703.uibk.ac.at
parents: 192
diff changeset
    85
  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
    86
  environment of Isabelle. It is therefore easy to experiment with what is 
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    87
  displayed. However, for better readability we will drop the enclosing 
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    88
  \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"} and just write:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    89
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    90
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    91
  @{ML [display,gray] "3 + 4"}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    92
  
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    93
  Whenever appropriate we also show the response the code 
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    94
  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
    95
  @{text [quotes] ">"}, like:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    96
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    97
  @{ML_response [display,gray] "3 + 4" "7"}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    98
195
7305beb69893 corrected typos + some small reformulations
griff@colo2-c703.uibk.ac.at
parents: 192
diff changeset
    99
  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
   100
  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
   101
  \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
   102
  command needs to be run in a UNIX-shell, for example:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   103
181
5baaabe1ab92 updated to new method_setup
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   104
  @{text [display] "$ grep -R ThyOutput *"}
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   105
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   106
  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
   107
  \textit{italic} and highlighted as follows:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   108
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   109
  \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
   110
  Further information or pointers to files.
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   111
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   112
182
4d0e2edd476d added hyperlinks for every file pointer
Christian Urban <urbanc@in.tum.de>
parents: 181
diff changeset
   113
  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
   114
  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
   115
  {http://isabelle.in.tum.de/repos/isabelle/}.
4d0e2edd476d added hyperlinks for every file pointer
Christian Urban <urbanc@in.tum.de>
parents: 181
diff changeset
   116
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
   117
  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
   118
  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
   119
  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
   120
*}
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   121
233
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   122
section {* Naming Conventions in the Isabelle Sources *}
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   123
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   124
text {*
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   125
  There are a few naming conventions in Isabelle that might aid reading 
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   126
  and writing code. (Remember that code is written once, but read numerous 
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   127
  times.) The most important conventions are:
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   128
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   129
  \begin{itemize}
234
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   130
  \item @{text t}, @{text u} for (raw) terms; ML-type: @{ML_type term}
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   131
  \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
   132
  \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
   133
  \item @{text th}, @{text thm} for theorems; ML-type: @{ML_type thm}
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   134
  \item @{text "foo_tac"} for tactics; ML-type: @{ML_type tactic}
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   135
  \item @{text thy} for theories; ML-type: @{ML_type theory}
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   136
  \item @{text ctxt} for proof contexts; ML-type: @{ML_type Proof.context}
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   137
  \item @{text lthy} for local theories; ML-type: @{ML_type local_theory}
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   138
  \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
   139
  \item @{text mx} for mixfix syntax annotations; ML-type @{ML_type mixfix}
233
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   140
  \end{itemize}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
   141
*}
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
   142
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   143
section {* Acknowledgements *}
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   144
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   145
text {*
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   146
  Financial support for this tutorial was provided by the German 
122
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   147
  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
   148
  people contributed to the text:
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   149
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   150
  \begin{itemize}
122
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   151
  \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
   152
  \simpleinductive-package and the code for the @{text
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   153
  "chunk"}-antiquotation. He also wrote the first version of the chapter
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   154
  describing the package and has been helpful \emph{beyond measure} with
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   155
  answering questions about Isabelle.
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   156
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   157
  \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
   158
  \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
   159
  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
   160
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   161
  \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
   162
  about parsing.
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   163
180
9c25418db6f0 added a recipy about SAT solvers
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   164
  \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
   165
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   166
  \item {\bf Alexander Krauss} wrote the first version of the ``first-steps'' 
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 137
diff changeset
   167
  chapter and also contributed the material on @{text NamedThmsFun}.
194
8cd51a25a7ca some polishing
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   168
207
d3cd633e8240 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 203
diff changeset
   169
  \item {\bf Christian Sternagel} proofread the tutorial and made 
d3cd633e8240 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 203
diff changeset
   170
  comments on the text. 
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   171
  \end{itemize}
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   172
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   173
  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
   174
  errors lies with me.\bigskip
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   175
233
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   176
  \vspace{5cm}
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   177
  {\Large\bf
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   178
  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
   179
  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
   180
  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
   181
  with TBD.}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   182
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   183
  
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   184
  \vfill
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   185
  This document was compiled with:\\
228
fe45fbb111c5 various additions
Christian Urban <urbanc@in.tum.de>
parents: 207
diff changeset
   186
  \input{version}\\
fe45fbb111c5 various additions
Christian Urban <urbanc@in.tum.de>
parents: 207
diff changeset
   187
  \input{pversion}
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   188
*}
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   190
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   191
195
7305beb69893 corrected typos + some small reformulations
griff@colo2-c703.uibk.ac.at
parents: 192
diff changeset
   192
end