CookBook/Intro.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 16 Mar 2009 00:12:32 +0100
changeset 178 fb8f22dd8ad0
parent 177 4e2341f6599d
child 180 9c25418db6f0
permissions -rw-r--r--
adapted to latest Attrib.setup changes and more work on the simple induct chapter
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
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
begin
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
64
9a6e5e0c4906 deleted old files and added code to give a special tag to the command ML
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
     6
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
chapter {* Introduction *}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
text {*
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    10
  If your next project requires you to program on the ML-level of Isabelle,
106
bdd82350cf22 renamed in the pdf all instances of cookbook to tutorial (in order to sound more serious)
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    11
  then this tutorial is for you. It will guide you through the first steps of
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    12
  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
    13
  get to know the ML-level of Isabelle is by experimenting with the many code
106
bdd82350cf22 renamed in the pdf all instances of cookbook to tutorial (in order to sound more serious)
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    14
  examples included in the tutorial. The code is as far as possible checked
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    15
  against recent versions of Isabelle.  If something does not work, then
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    16
  please let us know. If you have comments, criticism or like to add to the
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
    17
  tutorial, please feel free---you are most welcome! The tutorial is meant to be 
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 122
diff changeset
    18
  gentle and comprehensive. To achieve this we need your feedback. 
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
*}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
section {* Intended Audience and Prior Knowledge *}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
text {* 
106
bdd82350cf22 renamed in the pdf all instances of cookbook to tutorial (in order to sound more serious)
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    24
  This tutorial targets readers who already know how to use Isabelle for
75
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
12
2f1736cb8f26 various changes by Alex and Christian
Christian Urban <urbanc@in.tum.de>
parents: 11
diff changeset
    56
  Then of course there is:
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    57
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    58
  \begin{description}
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  \item[The code] is of course the ultimate reference for how
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
  things really work. Therefore you should not hesitate to look at the
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
  way things are actually implemented. More importantly, it is often
89
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
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
    63
  to learn from that code.
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
  \end{description}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
*}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    68
section {* Typographic Conventions *}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
    69
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
    70
text {*
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    71
106
bdd82350cf22 renamed in the pdf all instances of cookbook to tutorial (in order to sound more serious)
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    72
  All ML-code in this tutorial is typeset in highlighted boxes, like the following 
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    73
  ML-expression:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    74
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    75
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    76
  \begin{graybox}
85
b02904872d6b better handling of {* and *}
Christian Urban <urbanc@in.tum.de>
parents: 84
diff changeset
    77
  \isacommand{ML}~@{text "\<verbopen>"}\isanewline
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    78
  \hspace{5mm}@{ML "3 + 4"}\isanewline
85
b02904872d6b better handling of {* and *}
Christian Urban <urbanc@in.tum.de>
parents: 84
diff changeset
    79
  @{text "\<verbclose>"}
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    80
  \end{graybox}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    81
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    82
  
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    83
  These boxes corresponds to how code can be processed inside the interactive
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    84
  environment of Isabelle. It is therefore easy to experiment with what is 
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    85
  displayed. However, for better readability we will drop the enclosing 
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    86
  \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"} and just write:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    87
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    88
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    89
  @{ML [display,gray] "3 + 4"}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    90
  
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    91
  Whenever appropriate we also show the response the code 
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    92
  generates when evaluated. This response is prefixed with a 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
    93
  @{text [quotes] ">"}, like:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    94
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    95
  @{ML_response [display,gray] "3 + 4" "7"}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    96
122
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
    97
  The user-level commands of Isabelle (i.e.~the non-ML code) are written
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
    98
  in bold, for example \isacommand{lemma}, \isacommand{apply},
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
    99
  \isacommand{foobar} and so on.  We use @{text "$"} to indicate that a
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   100
  command needs to be run in a Unix-shell, for example:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   101
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   102
  @{text [display] "$ ls -la"}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   103
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   104
  Pointers to further information and Isabelle files are typeset in 
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   105
  italic and highlighted as follows:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   106
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   107
  \begin{readmore}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   108
  Further information or pointers to files.
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   109
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   110
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
   111
  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
   112
  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
   113
  to solve the exercises on your own, and then look at the solutions.
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   114
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
   115
*}
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
   116
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   117
section {* Acknowledgements *}
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   118
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   119
text {*
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   120
  Financial support for this tutorial was provided by the German 
122
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   121
  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
   122
  people contributed to the text:
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   123
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   124
  \begin{itemize}
122
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   125
  \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
   126
  \simpleinductive-package and the code for the @{text
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   127
  "chunk"}-antiquotation. He also wrote the first version of the chapter
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   128
  describing the package and has been helpful \emph{beyond measure} with
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   129
  answering questions about Isabelle.
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   130
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   131
  \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
   132
  \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
   133
  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
   134
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   135
  \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
   136
  about parsing.
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   137
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   138
  \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
   139
  chapter and also contributed the material on @{text NamedThmsFun}.
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   140
  \end{itemize}
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   141
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   142
  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
   143
  errors lies with me.\bigskip
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   144
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   145
  {\Large\bf
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   146
  This document is still in the process of being written! All of the
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   147
  text is still under constructions. Sections and 
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   148
  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
   149
  with TBD.}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   150
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   151
  
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   152
  \vfill
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   153
  This document was compiled with:\\
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   154
  \input{version}
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   155
*}
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   157
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   158
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
end