ProgTutorial/Intro.thy
author Norbert Schirmer <norbert.schirmer@web.de>
Wed, 22 May 2019 13:24:30 +0200
changeset 575 c3dbc04471a9
parent 574 034150db9d91
child 578 69c78980c8a4
permissions -rw-r--r--
fixing some Line references
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
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
     5
chapter \<open>Introduction\<close>
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
     7
text \<open>
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
     8
   \begin{flushright}
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
     9
  {\em
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
    10
  ``My thesis is that programming is not at the bottom of the intellectual \\
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
    11
  pyramid, but at the top. It's creative design of the highest order. It \\
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
    12
  isn't monkey or donkey work; rather, as Edsger Dijkstra famously \\
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
    13
  claimed, it's amongst the hardest intellectual tasks ever attempted.''} \\[1ex]
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 306
diff changeset
    14
  Richard Bornat, In {\em Defence of Programming}. \cite{Bornat-lecture}
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
    15
  \end{flushright}
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
    16
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
    17
  \medskip
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    18
  If your next project requires you to program Isabelle with ML,
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
    19
  then this tutorial is for you. It will guide you through the first steps of
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 435
diff changeset
    20
  Isabelle programming, and also explain ``tricks of the trade''. We also hope
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 435
diff changeset
    21
  the tutorial will encourage students and researchers to play with Isabelle
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 435
diff changeset
    22
  and implement new ideas. The source code of Isabelle can look intimidating,
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
    23
  but beginners can get by with knowledge of only a handful of concepts, 
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
    24
  a small number of functions and a few basic coding conventions.
557
77ea2de0ca62 updated for Isabelle 2014
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 553
diff changeset
    25
  There is also a considerable amount of code written in Scala that allows
77ea2de0ca62 updated for Isabelle 2014
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 553
diff changeset
    26
  Isabelle interface with the Jedit GUI. Explanation of this part is beyond
77ea2de0ca62 updated for Isabelle 2014
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 553
diff changeset
    27
  this tutorial.
356
43df2d59fb98 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 353
diff changeset
    28
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    29
  The best way to get to know the Isabelle/ML is by experimenting
356
43df2d59fb98 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 353
diff changeset
    30
  with the many code examples included in the tutorial. The code is as far as
43df2d59fb98 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 353
diff changeset
    31
  possible checked against the Isabelle
540
d144fc51fe04 removed versions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 536
diff changeset
    32
  distribution.%%\footnote{\input{version.tex}}
535
5734ab5dd86d adapted to new build framework
Christian Urban <urbanc@in.tum.de>
parents: 532
diff changeset
    33
  If something does not work, then
356
43df2d59fb98 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 353
diff changeset
    34
  please let us know. It is impossible for us to know every environment,
43df2d59fb98 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 353
diff changeset
    35
  operating system or editor in which Isabelle is used. If you have comments,
43df2d59fb98 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 353
diff changeset
    36
  criticism or like to add to the tutorial, please feel free---you are most
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 435
diff changeset
    37
  welcome!! The tutorial is meant to be gentle and comprehensive. To achieve
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 392
diff changeset
    38
  this we need your help and feedback.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
    39
\<close>
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
    41
section \<open>Intended Audience and Prior Knowledge\<close>
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
    43
text \<open>
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
    44
  This tutorial targets readers who already know how to use Isabelle for
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    45
  writing theories and proofs. We also assume that readers are familiar with
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    46
  the functional programming language ML, the language in which most of
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    47
  Isabelle is implemented. If you are unfamiliar with either of these two
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    48
  subjects, then you should first work through the Isabelle/HOL tutorial
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    49
  @{cite "isa-tutorial"} or Paulson's book on ML @{cite "paulson-ml2"}. Recently,
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
    50
  Isabelle has adopted a sizable amount of Scala code for a slick GUI
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
    51
  based on jEdit. This part of the code is beyond the interest of this
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
    52
  tutorial, since it mostly does not concern the regular Isabelle 
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
    53
  developer.
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    54
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    55
  The rich Isabelle infrastructure can be categorized by various aspects @{cite "wenzel-technology"}:
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    56
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    57
     @{emph \<open>@{bold \<open>Logic\<close>}\<close>} 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    58
      \begin{description}
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    59
        \item[Isabelle/Pure] is the logical Framework and bootstrap environment. The Pure logic
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    60
          is used to represent rules for Higher-Order Natural Deduction declaratively. This allows
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    61
          the implementation and definition of object logics like HOL using the Pure logic and
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    62
          framework.
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    63
        \item[Isabelle/HOL] is the main library of theories and tools for applications that is used 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    64
          throughout this tutorial.
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    65
      \end{description}
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    66
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    67
     @{emph \<open>@{bold\<open>Programming\<close>}\<close>} 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    68
      \begin{description}
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    69
        \item[Isabelle/ML] is the Isabelle tool implementation and extension language. It is based
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    70
          on Poly/ML\footnote{@{url \<open>http://polyml.org\<close>}}. Both Isabelle/Pure and Isabelle/ML emerge 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    71
          from the same bootstrap process: the result is a meta-language for programming the logic
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    72
          that is intertwined with it from a technological viewpoint, but logic and programming
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    73
          remain formally separated.
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    74
        \item[Isabelle/Scala] is the Isabelle system programming language. It connects the logical
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    75
          environment with the outside world. Most notably resulting in the Prover IDE Isabelle/jEdit 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    76
          and the command line tools.
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    77
      \end{description}
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    78
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    79
     @{emph \<open>@{bold\<open>Proof\<close>}\<close>} 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    80
      \begin{description}
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    81
        \item[Isabelle/Isar] is the structured proof language of the Isabelle framework. Isar 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    82
          means, Intelligible semi-automated reasoning. 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    83
        \item[Document language] for HTML output and \LaTeX type-setting of proof text. A proof
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    84
          document combines formal and informal text to describe what has been proven to a general
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    85
          audience.
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    86
      \end{description}
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    87
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    88
     @{emph \<open>@{bold\<open>IDE\<close>}\<close>} 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    89
      \begin{description}
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    90
        \item[Isabelle/jEdit] is the IDE for proof and tool development. It provides a rich interactive
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    91
          frontend to the Isabelle framework in which logic and proof development, document creation as
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    92
          well as ML programming are seamlessly integrated. 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    93
      \end{description}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
    94
\<close>
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
    96
section \<open>Existing Documentation\<close>
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
    98
text \<open>
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
    99
  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
   100
  part of the distribution of Isabelle):
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
  \begin{description}
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   103
  \item[The Isabelle/Isar Reference Manual] provides a top level view on the Isabelle system,
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   104
  explaining general concepts and specification material (like grammars,
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   105
  examples and so on) about Isabelle, Isar, Pure, HOL and the document language.
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   107
  \item[The Isabelle/Isar Implementation Manual] describes Isabelle
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   108
  implementation from a high-level perspective, documenting the major 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   109
  underlying concepts and interfaces. 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   110
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   111
  \item[Isabelle/jEdit] describes the IDE.
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   112
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   113
  \item[The Old Introduction to Isabelle] 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
   114
  to be the main reference of Isabelle at a time when all proof scripts 
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   115
  were written with ML. 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
   116
  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
   117
  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
   118
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   119
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
   120
  \end{description}
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
234
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   122
  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
   123
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
   124
  \begin{description}
234
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   125
  \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
   126
  things really work. Therefore you should not hesitate to look at the
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   127
  way things are actually implemented. While much of the Isabelle
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   128
  code is uncommented, some parts have very helpful comments---particularly
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   129
  the code about theorems and terms. Despite the lack of comments in most
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   130
  parts, it is often good to look at code that does similar things as you 
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   131
  want to do and learn from it. 
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   132
  This tutorial contains frequently pointers to the
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   133
  Isabelle sources. The best way is to interactively explore the sources
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   134
  within the IDE provided by Isabelle/jEdit. By loading @{ML_file "Pure/ROOT.ML"}
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   135
  into Isabelle/jEdit the sources of Pure are annotated with markup and you can
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   136
  interactively follow the structure.
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   137
  Moreover, the UNIX command \mbox{\<open>grep -R\<close>} or hypersearch within Isabelle/jEdit is
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   138
  often your best friend while programming with Isabelle. To understand the sources,
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   139
  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
   140
  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
   141
  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
   142
  files and ``change sets''.
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
  \end{description}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   144
\<close>
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   146
section \<open>Typographic Conventions\<close>
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
   147
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   148
text \<open>
181
5baaabe1ab92 updated to new method_setup
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   149
  All ML-code in this tutorial is typeset in shaded boxes, like the following 
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   150
  simple ML-expression:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   151
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   152
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   153
  \begin{graybox}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   154
  \isacommand{ML}~\<open>\<verbopen>\<close>\isanewline
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 568
diff changeset
   155
  \hspace{5mm}@{ML \<open>3 + 4\<close>}\isanewline
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   156
  \<open>\<verbclose>\<close>
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   157
  \end{graybox}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   158
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   159
  
195
7305beb69893 corrected typos + some small reformulations
griff@colo2-c703.uibk.ac.at
parents: 192
diff changeset
   160
  These boxes correspond to how code can be processed inside the interactive
343
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 335
diff changeset
   161
  environment of Isabelle. It is therefore easy to experiment with the code
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 435
diff changeset
   162
  that is shown in this tutorial. However, for better readability we will drop
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   163
  the enclosing \isacommand{ML}~\<open>\<verbopen> \<dots> \<verbclose>\<close> and just
343
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 335
diff changeset
   164
  write:
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
   165
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 568
diff changeset
   166
  @{ML [display,gray] \<open>3 + 4\<close>}
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   167
  
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   168
  Whenever appropriate we also show the response the code 
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
   169
  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
   170
  @{text [quotes] ">"}, like:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   171
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 568
diff changeset
   172
  @{ML_matchresult [display,gray] \<open>3 + 4\<close> \<open>7\<close>}
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   173
195
7305beb69893 corrected typos + some small reformulations
griff@colo2-c703.uibk.ac.at
parents: 192
diff changeset
   174
  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
   175
  in \isacommand{bold face} (e.g., \isacommand{lemma}, \isacommand{apply},
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   176
  \isacommand{foobar} and so on).  We use \<open>$ \<dots>\<close> to indicate that a
234
f84bc59cb5be polished
Christian Urban <urbanc@in.tum.de>
parents: 233
diff changeset
   177
  command needs to be run in a UNIX-shell, for example:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   178
441
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
   179
  @{text [display] "$ grep -R Thy_Output *"}
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   180
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   181
  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
   182
  \textit{italic} and highlighted as follows:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   183
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   184
  \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
   185
  Further information or pointers to files.
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   186
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   187
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 435
diff changeset
   188
  Note that pointers to Isabelle files are hyperlinked to the tip of the Mercurial
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 392
diff changeset
   189
  repository at \href{http://isabelle.in.tum.de/repos/isabelle/}
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 435
diff changeset
   190
  {http://isabelle.in.tum.de/repos/isabelle/}, not the latest stable release
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 392
diff changeset
   191
  of Isabelle.
182
4d0e2edd476d added hyperlinks for every file pointer
Christian Urban <urbanc@in.tum.de>
parents: 181
diff changeset
   192
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 167
diff changeset
   193
  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
   194
  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
   195
  to solve the exercises on your own, and then look at the solutions.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   196
\<close>
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   197
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   198
section \<open>How To Understand Isabelle Code\<close>
421
620a24bf954a added a section to the introduction; described @{make_string}
Christian Urban <urbanc@in.tum.de>
parents: 417
diff changeset
   199
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   200
text \<open>
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   201
  One of the more difficult aspects of any kind of programming is to
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   202
  understand code written by somebody else. This is aggravated in Isabelle by
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   203
  the fact that many parts of the code contain none or only few
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   204
  comments. There is one strategy that might be helpful to navigate your way:
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   205
  ML and Isabelle/jEdit is an interactive programming environment, which means you can evaluate
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   206
  code on the fly (for example inside an \isacommand{ML}~\<open>\<verbopen>\<dots>\<verbclose>\<close> section). So you can copy (self-contained)
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   207
  chunks of existing code into a separate theory file and then study it
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   208
  alongside with examples. You can also install ``probes'' inside the copied
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   209
  code without having to recompile the whole Isabelle distribution. Such
421
620a24bf954a added a section to the introduction; described @{make_string}
Christian Urban <urbanc@in.tum.de>
parents: 417
diff changeset
   210
  probes might be messages or printouts of variables (see chapter
568
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   211
  \ref{chp:firststeps}). Moreover, PolyML contains a debugger that is also supported from Isabelle/jEdit
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   212
  with breakpoints and stack inspection.
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   213
  Still it seems that
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   214
  probing the code with explicit print statements is an effective method
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   215
  for understanding what some piece of code is doing. However do not expect
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   216
  quick results with this! It is painful. Depending on the size of the code
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   217
  you are looking at, you will spend the better part of a quiet afternoon with
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   218
  it. And there seems to be no better way for understanding code in Isabelle.
568
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   219
  Summarizing: 
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   220
  \begin{itemize}
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   221
     \item Get familiar with the concepts and the applications by studying the Isabelle/Isar 
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   222
     reference manual and the rich set of theories and libraries in the standard distribution and the AFP.
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   223
     \item Get familiar with the implementation and the best practices (like coding style, 
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   224
     canonical argument ordering, \<open>\<dots>\<close>) by reading the Isabelle/Isar implementation manual and by 
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   225
     browsing, reading and digesting both the code implementing a function as well and the 
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   226
     applications of that function.
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   227
     \item Interactively experiment with the code.
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   228
  \end{itemize}
be23597e81db adding to "how to understand code"
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   229
  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   230
\<close>
421
620a24bf954a added a section to the introduction; described @{make_string}
Christian Urban <urbanc@in.tum.de>
parents: 417
diff changeset
   231
620a24bf954a added a section to the introduction; described @{make_string}
Christian Urban <urbanc@in.tum.de>
parents: 417
diff changeset
   232
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   233
section \<open>Aaaaargh! My Code Does not Work Anymore\<close>
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   234
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   235
text \<open>
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   236
  One unpleasant aspect of any code development inside a larger system is that
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 435
diff changeset
   237
  one has to aim at a ``moving target''. Isabelle is no exception of this. Every 
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   238
  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
   239
  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
   240
  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
   241
  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
   242
  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
   243
  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
   244
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   245
  However, there are a few steps you can take to mitigate unwanted
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   246
  interferences with code changes from other developers. First, you can base
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   247
  your code on the latest stable release of Isabelle (it is aimed to have one
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   248
  such release at least once every year). This might cut you off from the
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   249
  latest feature implemented in Isabelle, but at least you do not have to
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   250
  track side-steps or dead-ends in the Isabelle development. Of course this
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   251
  means also you have to synchronise your code at the next stable release. If
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   252
  you do not synchronise, be warned that code seems to ``rot'' very
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   253
  quickly. Another possibility is to get your code into the Isabelle
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   254
  distribution. For this you have to convince other developers that your code
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   255
  or project is of general interest. If you managed to do this, then the
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   256
  problem of the moving target goes away, because when checking in new code,
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   257
  developers are strongly urged to test it against Isabelle's code base. If
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   258
  your project is part of that code base, then maintenance is done by
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   259
  others. Unfortunately, this might not be a helpful advice for all types of
454
e2fe7e93333c corrected AFP name
griff
parents: 453
diff changeset
   260
  projects. A lower threshold for inclusion has the Archive of Formal
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   261
  Proofs, short AFP.\footnote{\url{http://afp.sourceforge.net/}} This archive
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   262
  has been created mainly for formalisations that are interesting but not
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   263
  necessarily of general interest. If you have ML-code as part of a
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   264
  formalisation, then this might be the right place for you. There is no
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   265
  problem with updating your code after submission. At the moment developers
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   266
  are not as diligent with checking their code against the AFP than with
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   267
  checking agains the distribution, but generally problems will be caught and
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   268
  the developer, who caused them, is expected to fix them. So also in this
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   269
  case code maintenance is done for you.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   270
\<close>
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   271
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   272
section \<open>Serious Isabelle ML-Programming\<close>
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   273
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   274
text \<open>
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   275
  As already pointed out in the previous section, Isabelle is a joint effort 
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   276
  of many developers. Therefore, disruptions that break the work of others
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   277
  are generally frowned upon. ``Accidents'' however do happen and everybody knows
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   278
  this. Still to keep them to a minimum, you can submit your changes first to a rather 
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   279
  sophisticated \emph{testboard}, which will perform checks of your changes against the
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   280
  Isabelle repository and against the AFP. The advantage of the testboard is
502
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
   281
  that the testing is performed by rather powerful machines saving you lengthy
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   282
  tests on, for example, your own laptop. You can see the results of the testboard 
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   283
  at 
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   284
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   285
  \begin{center}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   286
  \url{http://isabelle.in.tum.de/testboard/Isabelle/}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   287
  \end{center}
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   288
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   289
  which is organised like a Mercurial repository. A green point next to a change
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   290
  indicates that the change passes the corresponding tests (for this of course you
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   291
  have to allow some time). You can summit any changes to the testboard using the 
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   292
  command
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   293
532
Christian Urban <urbanc@in.tum.de>
parents: 530
diff changeset
   294
  @{text [display] "$ hg push -f ssh://...@hgbroy.informatik.tu-muenchen.de\\
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   295
   //home/isabelle-repository/repos/testboard"}
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   296
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   297
  where the dots need to be replaced by your login name.  Note that for
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   298
  pushing changes to the testboard you need to add the option \<open>-f\<close>,
502
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
   299
  which should \emph{never} be used with the main Isabelle
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   300
  repository. While the testboard is a great system for supporting Isabelle
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   301
  developers, its disadvantage is that it needs login permissions for the
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   302
  computers in Munich. So in order to use it, you might have to ask other
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   303
  developers to obtain one.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   304
\<close>
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   305
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 469
diff changeset
   306
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   307
section \<open>Some Naming Conventions in the Isabelle Sources\<close>
233
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   308
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   309
text \<open>
254
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
   310
  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
   311
  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
   312
  times.) The most important conventions are:
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   313
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   314
  \begin{itemize}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   315
  \item \<open>t\<close>, \<open>u\<close>, \<open>trm\<close> for (raw) terms; ML-type: @{ML_type term}
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   316
  \item \<open>ct\<close>, \<open>cu\<close> for certified terms; ML-type: @{ML_type cterm}
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   317
  \item \<open>ty\<close>, \<open>T\<close>, \<open>U\<close> for (raw) types; ML-type: @{ML_type typ}
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   318
  \item \<open>S\<close> for sorts; ML-type: @{ML_type sort}
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   319
  \item \<open>th\<close>, \<open>thm\<close> for theorems; ML-type: @{ML_type thm}
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   320
  \item \<open>foo_tac\<close> for tactics; ML-type: @{ML_type tactic}
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   321
  \item \<open>thy\<close> for theories; ML-type: @{ML_type theory}
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   322
  \item \<open>ctxt\<close> for proof contexts; ML-type: @{ML_type Proof.context}
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   323
  \item \<open>lthy\<close> for local theories; ML-type: @{ML_type local_theory}
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   324
  \item \<open>context\<close> for generic contexts; ML-type @{ML_type Context.generic}
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   325
  \item \<open>mx\<close> for mixfix syntax annotations; ML-type @{ML_type mixfix}
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   326
  \item \<open>prt\<close> for pretty printing; ML-type @{ML_type Pretty.T}
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   327
  \item \<open>phi\<close> for morphisms; ML-type @{ML_type morphism}
233
61085dd44e8c added a section about naming conventions
Christian Urban <urbanc@in.tum.de>
parents: 228
diff changeset
   328
  \end{itemize}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   329
\<close>
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
   330
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   331
section \<open>Acknowledgements\<close>
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   332
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   333
text \<open>
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   334
  Financial support for this tutorial was provided by the German 
122
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   335
  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
   336
  people contributed to the text:
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   337
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   338
  \begin{itemize}
122
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 121
diff changeset
   339
  \item {\bf Stefan Berghofer} wrote nearly all of the ML-code of the
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   340
  \simpleinductive-package and the code for the \<open>chunk\<close>-antiquotation. He also wrote the first version of chapter
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 392
diff changeset
   341
  \ref{chp:package} describing this package and has been helpful \emph{beyond
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 392
diff changeset
   342
  measure} with answering questions about Isabelle.
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   343
469
7a558c5119b2 added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents: 466
diff changeset
   344
  \item {\bf Jasmin Blanchette} helped greatly with section \ref{sec:pretty}
7a558c5119b2 added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents: 466
diff changeset
   345
  and exercise \ref{fun:killqnt}.
248
11851b20fb78 added more to the pretty section and updated the acknowledgements
Christian Urban <urbanc@in.tum.de>
parents: 235
diff changeset
   346
574
034150db9d91 polish document
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   347
  \item {\bf Sascha B\"ohme} contributed the recipes in \ref{rec:timeout}, 
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   348
  \ref{rec:external} and \ref{rec:oracle}. He also wrote section \ref{sec:conversion} 
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   349
  and helped with recipe \ref{rec:timing}. Parts of section \ref{sec:storing}
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   350
  are by him.
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   351
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 392
diff changeset
   352
  \item {\bf Lukas Bulwahn} made me aware of a problem with recursive
530
Christian Urban <urbanc@in.tum.de>
parents: 529
diff changeset
   353
  parsers, contributed exercise \ref{ex:contextfree} and contributed 
Christian Urban <urbanc@in.tum.de>
parents: 529
diff changeset
   354
  to the ``introspection'' of theorems in section \ref{sec:theorems}.
529
13d7ea419c5f moved the introspection part into the theorem section
Christian Urban <urbanc@in.tum.de>
parents: 526
diff changeset
   355
390
8ad407e77ea0 added example by Lukas Bulwahn
Christian Urban <urbanc@in.tum.de>
parents: 358
diff changeset
   356
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 392
diff changeset
   357
  \item {\bf Jeremy Dawson} wrote the first version of chapter \ref{chp:parsing}
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   358
  about parsing.
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   359
553
c53d74b34123 updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 542
diff changeset
   360
  %%\item {\bf Florian Haftmann} helped with maintaining recipe \ref{rec:callml}.
542
4b96e3c8b33e updated the CallML section with the help from Florian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 540
diff changeset
   361
180
9c25418db6f0 added a recipy about SAT solvers
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   362
  \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
   363
530
Christian Urban <urbanc@in.tum.de>
parents: 529
diff changeset
   364
  \item {\bf Rafal Kolanski} contributed to the ``introspection'' of theorems 
Christian Urban <urbanc@in.tum.de>
parents: 529
diff changeset
   365
  in section \ref{sec:theorems}.
526
9e191bc4a828 polished
Christian Urban <urbanc@in.tum.de>
parents: 525
diff changeset
   366
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 392
diff changeset
   367
  \item {\bf Alexander Krauss} wrote a very early version of the ``first-steps'' 
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   368
  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
   369
553
c53d74b34123 updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 542
diff changeset
   370
  %%\item {\bf Tobias Nipkow} contributed recipe \ref{rec:callml}.
427
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents: 425
diff changeset
   371
417
5f00958e3c7b typos fixed by Michael Norrish
Christian Urban <urbanc@in.tum.de>
parents: 414
diff changeset
   372
  \item {\bf Michael Norrish} proofread parts of the text.
5f00958e3c7b typos fixed by Michael Norrish
Christian Urban <urbanc@in.tum.de>
parents: 414
diff changeset
   373
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   374
  \item {\bf Norbert Schirmer} updated the document to work with Isabelle 2018 and 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   375
  later.
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   376
435
524b72520c43 added Andreas
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
   377
  \item {\bf Andreas Schropp} improved and corrected section \ref{sec:univ} and
524b72520c43 added Andreas
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
   378
   contributed towards section \ref{sec:sorts}.
524b72520c43 added Andreas
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
   379
207
d3cd633e8240 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 203
diff changeset
   380
  \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
   381
  many improvemets to the text. 
522
0ed6f49277c4 updated
Christian Urban <urbanc@in.tum.de>
parents: 521
diff changeset
   382
0ed6f49277c4 updated
Christian Urban <urbanc@in.tum.de>
parents: 521
diff changeset
   383
  \item {\bf Dmitriy Traytel} suggested to use the ML-antiquotation 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   384
  \<open>command_spec\<close> in section~\ref{sec:newcommand}, which simplified the code. 
560
8d30446d89f0 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   385
8d30446d89f0 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   386
  \item {\bf Piotr Trojanek} proofread the text.
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   387
  \end{itemize}
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   388
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   389
  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
   390
  errors lies with me.\bigskip
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   391
421
620a24bf954a added a section to the introduction; described @{make_string}
Christian Urban <urbanc@in.tum.de>
parents: 417
diff changeset
   392
  \newpage
620a24bf954a added a section to the introduction; described @{make_string}
Christian Urban <urbanc@in.tum.de>
parents: 417
diff changeset
   393
  \mbox{}\\[5cm]
620a24bf954a added a section to the introduction; described @{make_string}
Christian Urban <urbanc@in.tum.de>
parents: 417
diff changeset
   394
  
620a24bf954a added a section to the introduction; described @{make_string}
Christian Urban <urbanc@in.tum.de>
parents: 417
diff changeset
   395
  
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   396
  {\Large\bf
335
163ac0662211 reorganised the certified terms section; tuned
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
   397
  This tutorial 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
   398
  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
   399
  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
   400
  with TBD.}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   401
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
   402
  \vfill
540
d144fc51fe04 removed versions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 536
diff changeset
   403
  %%This document (version \input{tip.tex}\hspace{-0.5ex}) was compiled with:\\
d144fc51fe04 removed versions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 536
diff changeset
   404
  %%\input{version.tex}\\
535
5734ab5dd86d adapted to new build framework
Christian Urban <urbanc@in.tum.de>
parents: 532
diff changeset
   405
  %% \input{pversion}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   406
\<close>
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
195
7305beb69893 corrected typos + some small reformulations
griff@colo2-c703.uibk.ac.at
parents: 192
diff changeset
   408
end