CookBook/Intro.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 09 Feb 2009 04:18:14 +0000
changeset 107 258ce361ba1b
parent 106 bdd82350cf22
child 108 8bea3f74889d
permissions -rw-r--r--
polished and more material in the tactic 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
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
    17
  tutorial, feel free---you are most welcome!
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
*}
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
section {* Intended Audience and Prior Knowledge *}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
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
    23
  This tutorial targets readers who already know how to use Isabelle for
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    24
  writing theories and proofs. We also assume that readers are familiar with
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    25
  the functional programming language ML, the language in which most of
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    26
  Isabelle is implemented. If you are unfamiliar with either of these two
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    27
  subjects, you should first work through the Isabelle/HOL tutorial
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    28
  \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
    29
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
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    32
section {* Existing Documentation *}
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
text {*
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  
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
    36
  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
    37
  part of the distribution of Isabelle):
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  \begin{description}
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    40
  \item[The 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
    41
  from a high-level perspective, documenting both the underlying
6
007e09485351 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 5
diff changeset
    42
  concepts and some of the interfaces. 
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    44
  \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
    45
  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
    46
  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
    47
  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
    48
  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
    49
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
  \item[The Isar Reference Manual] is also an older document that provides
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
    51
  material about Isar and its implementation. Some material in it
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
    52
  is still useful.
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    53
  \end{description}
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
12
2f1736cb8f26 various changes by Alex and Christian
Christian Urban <urbanc@in.tum.de>
parents: 11
diff changeset
    55
  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
    56
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    57
  \begin{description}
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
  \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
    59
  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
    60
  way things are actually implemented. More importantly, it is often
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    61
  good to look at code that does similar things as you want to do and
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    62
  to learn from other people's code.
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  \end{description}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
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
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    67
section {* Typographic Conventions *}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
    68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
    69
text {*
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    70
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
    71
  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
    72
  ML-expression:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    73
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    74
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    75
  \begin{graybox}
85
b02904872d6b better handling of {* and *}
Christian Urban <urbanc@in.tum.de>
parents: 84
diff changeset
    76
  \isacommand{ML}~@{text "\<verbopen>"}\isanewline
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    77
  \hspace{5mm}@{ML "3 + 4"}\isanewline
85
b02904872d6b better handling of {* and *}
Christian Urban <urbanc@in.tum.de>
parents: 84
diff changeset
    78
  @{text "\<verbclose>"}
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    79
  \end{graybox}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    80
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    81
  
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    82
  These boxes corresponds to how code can be processed inside the interactive
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    83
  environment of Isabelle. It is therefore easy to experiment with what is 
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    84
  displayed. However, for better readability we will drop the enclosing 
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    85
  \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"} and just write:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    86
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    87
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    88
  @{ML [display,gray] "3 + 4"}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    89
  
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    90
  Whenever appropriate we also show the response the code 
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 80
diff changeset
    91
  generates when evaluated. This response is prefixed with a 
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
    92
  @{text [quotes] ">"} like:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    93
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    94
  @{ML_response [display,gray] "3 + 4" "7"}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
    95
101
Christian Urban <urbanc@in.tum.de>
parents: 89
diff changeset
    96
  The usual user-level commands of Isabelle are written in bold, for 
Christian Urban <urbanc@in.tum.de>
parents: 89
diff changeset
    97
  example \isacommand{lemma}, \isacommand{foobar} and so on.  
Christian Urban <urbanc@in.tum.de>
parents: 89
diff changeset
    98
  We use @{text "$"} to indicate that a command  needs to be run 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 106
diff changeset
    99
  in a Unix-shell, for example:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   100
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   101
  @{text [display] "$ ls -la"}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   102
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   103
  Pointers to further information and Isabelle files are typeset in 
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   104
  italic and highlighted as follows:
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   105
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   106
  \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
   107
  Further information or pointers to files.
75
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   108
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   109
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
   110
*}
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 66
diff changeset
   111
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
end