doc-src/IsarImplementation/Thy/ML.thy
author wenzelm
Fri, 22 Oct 2010 19:03:31 +0100
changeset 39882 ab0afd03a042
parent 39881 40aee0b95ddf
child 39883 3d3d6038bdaa
permissions -rw-r--r--
cover @{Isar.state};
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29755
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 29646
diff changeset
     1
theory "ML"
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 29646
diff changeset
     2
imports Base
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 29646
diff changeset
     3
begin
18538
88fe84d4d151 outline;
wenzelm
parents: 18537
diff changeset
     4
39822
0de42180febe basic setup for Chapter 0: Isabelle/ML;
wenzelm
parents: 36164
diff changeset
     5
chapter {* Isabelle/ML *}
20489
a684fc70d04e tentative appendix C;
wenzelm
parents: 18554
diff changeset
     6
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
     7
text {* Isabelle/ML is best understood as a certain culture based on
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
     8
  Standard ML.  Thus it is not a new programming language, but a
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
     9
  certain way to use SML at an advanced level within the Isabelle
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    10
  environment.  This covers a variety of aspects that are geared
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    11
  towards an efficient and robust platform for applications of formal
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    12
  logic with fully foundational proof construction --- according to
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    13
  the well-known \emph{LCF principle}.  There are specific library
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    14
  modules and infrastructure to address the needs for such difficult
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    15
  tasks.  For example, the raw parallel programming model of Poly/ML
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    16
  is presented as considerably more abstract concept of \emph{future
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    17
  values}, which is then used to augment the inference kernel, proof
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    18
  interpreter, and theory loader accordingly.
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    19
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    20
  The main aspects of Isabelle/ML are introduced below.  These
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    21
  first-hand explanations should help to understand how proper
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    22
  Isabelle/ML is to be read and written, and to get access to the
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    23
  wealth of experience that is expressed in the source text and its
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    24
  history of changes.\footnote{See
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    25
  \url{http://isabelle.in.tum.de/repos/isabelle} for the full
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    26
  Mercurial history.  There are symbolic tags to refer to official
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    27
  Isabelle releases, as opposed to arbitrary \emph{tip} versions that
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    28
  merely reflect snapshots that are never really up-to-date.}  *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    29
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    30
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    31
section {* Style and orthography *}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    32
39879
wenzelm
parents: 39878
diff changeset
    33
text {* The sources of Isabelle/Isar are optimized for
wenzelm
parents: 39878
diff changeset
    34
  \emph{readability} and \emph{maintainability}.  The main purpose is
wenzelm
parents: 39878
diff changeset
    35
  to tell an informed reader what is really going on and how things
wenzelm
parents: 39878
diff changeset
    36
  really work.  This is a non-trivial aim, but it is supported by a
wenzelm
parents: 39878
diff changeset
    37
  certain style of writing Isabelle/ML that has emerged from long
wenzelm
parents: 39878
diff changeset
    38
  years of system development.\footnote{See also the interesting style
wenzelm
parents: 39878
diff changeset
    39
  guide for OCaml
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    40
  \url{http://caml.inria.fr/resources/doc/guides/guidelines.en.html}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    41
  which shares many of our means and ends.}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    42
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    43
  The main principle behind any coding style is \emph{consistency}.
39879
wenzelm
parents: 39878
diff changeset
    44
  For a single author of a small program this merely means ``choose
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    45
  your style and stick to it''.  A complex project like Isabelle, with
39879
wenzelm
parents: 39878
diff changeset
    46
  long years of development and different contributors, requires more
wenzelm
parents: 39878
diff changeset
    47
  standardization.  A coding style that is changed every few years or
wenzelm
parents: 39878
diff changeset
    48
  with every new contributor is no style at all, because consistency
wenzelm
parents: 39878
diff changeset
    49
  is quickly lost.  Global consistency is hard to achieve, though.
wenzelm
parents: 39878
diff changeset
    50
  One should always strife at least for local consistency of modules
wenzelm
parents: 39878
diff changeset
    51
  and sub-systems, without deviating from some general principles how
wenzelm
parents: 39878
diff changeset
    52
  to write Isabelle/ML.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    53
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    54
  In a sense, a common coding style is like an \emph{orthography} for
39879
wenzelm
parents: 39878
diff changeset
    55
  the sources: it helps to read quickly over the text and see through
wenzelm
parents: 39878
diff changeset
    56
  the main points, without getting distracted by accidental
wenzelm
parents: 39878
diff changeset
    57
  presentation of free-style source.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    58
*}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    59
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    60
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    61
subsection {* Header and sectioning *}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    62
39879
wenzelm
parents: 39878
diff changeset
    63
text {* Isabelle source files have a certain standardized header
wenzelm
parents: 39878
diff changeset
    64
  format (with precise spacing) that follows ancient traditions
wenzelm
parents: 39878
diff changeset
    65
  reaching back to the earliest versions of the system by Larry
wenzelm
parents: 39878
diff changeset
    66
  Paulson.  E.g.\ see @{"file" "~~/src/Pure/thm.ML"}.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    67
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    68
  The header includes at least @{verbatim Title} and @{verbatim
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    69
  Author} entries, followed by a prose description of the purpose of
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    70
  the module.  The latter can range from a single line to several
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    71
  paragraphs of explanations.
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    72
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    73
  The rest of the file is divided into sections, subsections,
39879
wenzelm
parents: 39878
diff changeset
    74
  subsubsections, paragraphs etc.\ using some a layout via ML comments
wenzelm
parents: 39878
diff changeset
    75
  as follows.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    76
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    77
\begin{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    78
(*** section ***)
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    79
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    80
(** subsection **)
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    81
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    82
(* subsubsection *)
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    83
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    84
(*short paragraph*)
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    85
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    86
(*
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    87
  long paragraph
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    88
  more text
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    89
*)
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    90
\end{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    91
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    92
  As in regular typography, there is some extra space \emph{before}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    93
  section headings that are adjacent to plain text (not other headings
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    94
  as in the example above).
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    95
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    96
  \medskip The precise wording of the prose text given in these
39879
wenzelm
parents: 39878
diff changeset
    97
  headings is chosen carefully to indicate the main theme of the
wenzelm
parents: 39878
diff changeset
    98
  subsequent formal ML text.
wenzelm
parents: 39878
diff changeset
    99
*}
wenzelm
parents: 39878
diff changeset
   100
wenzelm
parents: 39878
diff changeset
   101
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   102
subsection {* Naming conventions *}
39879
wenzelm
parents: 39878
diff changeset
   103
wenzelm
parents: 39878
diff changeset
   104
text {* Since ML is the primary medium to express the meaning of the
wenzelm
parents: 39878
diff changeset
   105
  source text, naming of ML entities requires special care.
wenzelm
parents: 39878
diff changeset
   106
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   107
  \paragraph{Notation.}  A name consists of 1--3 \emph{words} (rarely
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   108
  4, but not more) that are separated by underscore.  There are three
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   109
  variants concerning upper or lower case letters, which are just for
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   110
  certain ML categories as follows:
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   111
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   112
  \medskip
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   113
  \begin{tabular}{lll}
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   114
  variant & example & ML categories \\\hline
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   115
  lower-case & @{verbatim foo_bar} & values, types, record fields \\
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   116
  capitalized & @{verbatim Foo_Bar} & datatype constructors, \\
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   117
    & & structures, functors \\
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   118
  upper-case & @{verbatim FOO_BAR} & special values (combinators), \\
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   119
    & & exception constructors, signatures \\
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   120
  \end{tabular}
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   121
  \medskip
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   122
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   123
  For historical reasons, many capitalized names omit underscores,
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   124
  e.g.\ old-style @{verbatim FooBar} instead of @{verbatim Foo_Bar}.
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   125
  Genuine mixed-case names are \emph{not} used --- clear division of
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   126
  words is essential for readability.\footnote{Camel-case was invented
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   127
  to workaround the lack of underscore in some early non-ASCII
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   128
  character sets and keywords.  Later it became a habitual in some
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   129
  language communities that are now strong in numbers.}
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   130
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   131
  A single (capital) character does not count as ``word'' in this
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   132
  respect: some Isabelle/ML are suffixed by extra markers like this:
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   133
  @{verbatim foo_barT}.
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   134
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   135
  Name variants are produced by adding 1--3 primes, e.g.\ @{verbatim
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   136
  foo'}, @{verbatim foo''}, or @{verbatim foo'''}, but not @{verbatim
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   137
  foo''''} or more.  Decimal digits scale better to larger numbers,
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   138
  e.g.\ @{verbatim foo0}, @{verbatim foo1}, @{verbatim foo42}.
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   139
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   140
  \paragraph{Scopes.}  Apart from very basic library modules, ML
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   141
  structures are not ``opened'', but names are referenced with
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   142
  explicit qualification, as in @{ML Syntax.string_of_term} for
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   143
  example.  When devising names for structures and their components it
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   144
  is important aim at eye-catching compositions of both parts, because
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   145
  this is how they are always seen in the sources and documentation.
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   146
  For the same reasons, aliases of well-known library functions should
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   147
  be avoided.
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   148
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   149
  Local names of function abstraction or case/lets bindings are
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   150
  typically shorter, sometimes using only rudiments of ``words'',
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   151
  while still avoiding cryptic shorthands.  An auxiliary function
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   152
  called @{verbatim helper}, @{verbatim aux}, or @{verbatim f} is
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   153
  considered bad style.
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   154
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   155
  Example:
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   156
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   157
  \begin{verbatim}
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   158
  (* RIGHT *)
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   159
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   160
  fun print_foo ctxt foo =
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   161
    let
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   162
      fun print t = ... Syntax.string_of_term ctxt t ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   163
    in ... end;
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   164
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   165
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   166
  (* RIGHT *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   167
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   168
  fun print_foo ctxt foo =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   169
    let
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   170
      val string_of_term = Syntax.string_of_term ctxt;
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   171
      fun print t = ... string_of_term t ...
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   172
    in ... end;
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   173
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   174
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   175
  (* WRONG *)
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   176
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   177
  val string_of_term = Syntax.string_of_term;
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   178
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   179
  fun print_foo ctxt foo =
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   180
    let
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   181
      fun aux t = ... string_of_term ctxt t ...
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   182
    in ... end;
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   183
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   184
  \end{verbatim}
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   185
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   186
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   187
  \paragraph{Specific conventions.} Here are some important specific
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   188
  name forms that occur frequently in the sources.
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   189
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   190
  \begin{itemize}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   191
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   192
  \item A function that maps @{verbatim foo} to @{verbatim bar} is
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   193
  called @{verbatim foo_to_bar} or @{verbatim bar_of_foo} (never
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   194
  @{verbatim foo2bar}, @{verbatim bar_from_foo}, @{verbatim
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   195
  bar_for_foo}, or @{verbatim bar4foo}).
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   196
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   197
  \item The name component @{verbatim legacy} means that the operation
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   198
  is about to be discontinued soon.
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   199
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   200
  \item The name component @{verbatim old} means that this is historic
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   201
  material that might disappear at some later stage.
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   202
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   203
  \item The name component @{verbatim global} means that this works
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   204
  with the background theory instead of the regular local context
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   205
  (\secref{sec:context}), sometimes for historical reasons, sometimes
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   206
  due a genuine lack of locality of the concept involved, sometimes as
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   207
  a fall-back for the lack of a proper context in the application
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   208
  code.  Whenever there is a non-global variant available, the
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   209
  application should be migrated to use it with a proper local
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   210
  context.
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   211
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   212
  \item Variables of the main context types of the Isabelle/Isar
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   213
  framework (\secref{sec:context} and \chref{ch:local-theory}) have
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   214
  firm naming conventions to allow to visualize the their data flow
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   215
  via plain regular expressions in the editor.  In particular:
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   216
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   217
  \begin{itemize}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   218
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   219
  \item theories are called @{verbatim thy}, rarely @{verbatim theory}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   220
  (never @{verbatim thry})
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   221
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   222
  \item proof contexts are called @{verbatim ctxt}, rarely @{verbatim
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   223
  context} (never @{verbatim ctx})
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   224
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   225
  \item generic contexts are called @{verbatim context}, rarely
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   226
  @{verbatim ctxt}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   227
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   228
  \item local theories are called @{verbatim lthy}, unless there is an
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   229
  implicit conversion to a general proof context (semantic super-type)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   230
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   231
  \end{itemize}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   232
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   233
  Variations with primed or decimal numbers are always possible, as
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   234
  well as ``sematic prefixes'' like @{verbatim foo_thy} or @{verbatim
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   235
  bar_ctxt}, but the base conventions above need to be preserved.
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   236
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   237
  \item The main logical entities (\secref{ch:logic}) also have
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   238
  established naming convention:
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   239
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   240
  \begin{itemize}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   241
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   242
  \item sorts are called @{verbatim S}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   243
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   244
  \item types are called @{verbatim T}, @{verbatim U}, or @{verbatim
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   245
  ty} (never @{verbatim t})
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   246
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   247
  \item terms are called @{verbatim t}, @{verbatim u}, or @{verbatim
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   248
  tm} (never @{verbatim trm})
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   249
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   250
  \item certified types are called @{verbatim cT}, rarely @{verbatim
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   251
  T}, with variants as for types
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   252
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   253
  \item certified terms are called @{verbatim ct}, rarely @{verbatim
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   254
  t}, with variants as for terms
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   255
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   256
  \item theorems are called @{verbatim th}, or @{verbatim thm}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   257
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   258
  \end{itemize}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   259
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   260
  Proper semantic names override these conventions completely.  For
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   261
  example, the left-hand side of an equation (as a term) can be called
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   262
  @{verbatim lhs} (not @{verbatim lhs_tm}).  Or a term that is treated
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   263
  specifically as a variable can be called @{verbatim v} or @{verbatim
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   264
  x}.
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   265
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   266
  \end{itemize}
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   267
*}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   268
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   269
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   270
subsection {* General source layout *}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   271
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   272
text {* The general Isabelle/ML source layout imitates regular
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   273
  type-setting to some extent, augmented by the requirements for
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   274
  deeply nested expressions that are commonplace in functional
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   275
  programming.
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   276
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   277
  \paragraph{Line length} is 80 characters according to ancient
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   278
  standards, but we allow as much as 100 characters, not
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   279
  more.\footnote{Readability requires to keep the beginning of a line
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   280
  in view while watching its end.  Modern wide-screen displays do not
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   281
  change the way how the human brain works.  As a rule of thumb,
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   282
  sources need to be printable on plain paper.} The extra 20
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   283
  characters acknowledge the space requirements due to qualified
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   284
  library references in Isabelle/ML.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   285
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   286
  \paragraph{White-space} is used to emphasize the structure of
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   287
  expressions, following mostly standard conventions for mathematical
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   288
  typesetting, as can be seen in plain {\TeX} or {\LaTeX}.  This
39879
wenzelm
parents: 39878
diff changeset
   289
  defines positioning of spaces for parentheses, punctuation, and
wenzelm
parents: 39878
diff changeset
   290
  infixes as illustrated here:
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   291
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   292
  \begin{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   293
  val x = y + z * (a + b);
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   294
  val pair = (a, b);
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   295
  val record = {foo = 1, bar = 2};
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   296
  \end{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   297
39879
wenzelm
parents: 39878
diff changeset
   298
  Lines are normally broken \emph{after} an infix operator or
wenzelm
parents: 39878
diff changeset
   299
  punctuation character.  For example:
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   300
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   301
  \begin{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   302
  val x =
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   303
    a +
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   304
    b +
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   305
    c;
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   306
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   307
  val tuple =
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   308
   (a,
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   309
    b,
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   310
    c);
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   311
  \end{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   312
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   313
  Some special infixes (e.g.\ @{verbatim "|>"}) work better at the
39879
wenzelm
parents: 39878
diff changeset
   314
  start of the line, but punctuation is always at the end.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   315
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   316
  Function application follows the tradition of @{text "\<lambda>"}-calculus,
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   317
  not informal mathematics.  For example: @{verbatim "f a b"} for a
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   318
  curried function, or @{verbatim "g (a, b)"} for a tupled function.
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   319
  Note that the space between @{verbatim g} and the pair @{verbatim
39879
wenzelm
parents: 39878
diff changeset
   320
  "(a, b)"} follows the important principle of
wenzelm
parents: 39878
diff changeset
   321
  \emph{compositionality}: the layout of @{verbatim "g p"} does not
wenzelm
parents: 39878
diff changeset
   322
  change when @{verbatim "p"} is refined to a concrete pair.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   323
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   324
  \paragraph{Indentation} uses plain spaces, never hard
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   325
  tabulators.\footnote{Tabulators were invented to move the carriage
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   326
  of a type-writer to certain predefined positions.  In software they
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   327
  could be used as a primitive run-length compression of consecutive
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   328
  spaces, but the precise result would depend on non-standardized
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   329
  editor configuration.}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   330
39879
wenzelm
parents: 39878
diff changeset
   331
  Each level of nesting is indented by 2 spaces, sometimes 1, very
wenzelm
parents: 39878
diff changeset
   332
  rarely 4, never 8.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   333
39879
wenzelm
parents: 39878
diff changeset
   334
  Indentation follows a simple logical format that only depends on the
wenzelm
parents: 39878
diff changeset
   335
  nesting depth, not the accidental length of the text that initiates
wenzelm
parents: 39878
diff changeset
   336
  a level of nesting.  Example:
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   337
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   338
  \begin{verbatim}
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   339
  (* RIGHT *)
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   340
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   341
  if b then
39879
wenzelm
parents: 39878
diff changeset
   342
    expr1_part1
wenzelm
parents: 39878
diff changeset
   343
    expr1_part2
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   344
  else
39879
wenzelm
parents: 39878
diff changeset
   345
    expr2_part1
wenzelm
parents: 39878
diff changeset
   346
    expr2_part2
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   347
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   348
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   349
  (* WRONG *)
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   350
39879
wenzelm
parents: 39878
diff changeset
   351
  if b then expr1_part1
wenzelm
parents: 39878
diff changeset
   352
            expr1_part2
wenzelm
parents: 39878
diff changeset
   353
  else expr2_part1
wenzelm
parents: 39878
diff changeset
   354
       expr2_part2
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   355
  \end{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   356
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   357
  The second form has many problems: it assumes a fixed-width font
39879
wenzelm
parents: 39878
diff changeset
   358
  when viewing the sources, it uses more space on the line and thus
wenzelm
parents: 39878
diff changeset
   359
  makes it hard to observe its strict length limit (working against
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   360
  \emph{readability}), it requires extra editing to adapt the layout
39879
wenzelm
parents: 39878
diff changeset
   361
  to changes of the initial text (working against
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   362
  \emph{maintainability}) etc.
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   363
39879
wenzelm
parents: 39878
diff changeset
   364
  \medskip For similar reasons, any kind of two-dimensional or tabular
wenzelm
parents: 39878
diff changeset
   365
  layouts, ASCII-art with lines or boxes of stars etc. should be
wenzelm
parents: 39878
diff changeset
   366
  avoided.
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   367
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   368
  \paragraph{Complex expressions} consisting of multi-clausal function
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   369
  definitions, @{verbatim case}, @{verbatim handle}, @{verbatim let},
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   370
  or combinations require special attention.  The syntax of Standard
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   371
  ML is a bit too ambitious in admitting too much variance that can
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   372
  distort the meaning of the text.
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   373
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   374
  Clauses of @{verbatim fun}, @{verbatim fn}, @{verbatim case},
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   375
  @{verbatim handle} get extra indentation to indicate the nesting
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   376
  clearly.  For example:
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   377
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   378
  \begin{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   379
  (* RIGHT *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   380
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   381
  fun foo p1 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   382
        expr1
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   383
    | foo p2 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   384
        expr2
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   385
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   386
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   387
  (* WRONG *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   388
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   389
  fun foo p1 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   390
    expr1
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   391
    | foo p2 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   392
    expr2
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   393
  \end{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   394
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   395
  Body expressions consisting of @{verbatim case} or @{verbatim let}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   396
  require care to maintain compositionality, to prevent loss of
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   397
  logical indentation where it is particularly imporant to see the
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   398
  structure of the text later on.  Example:
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   399
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   400
  \begin{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   401
  (* RIGHT *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   402
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   403
  fun foo p1 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   404
        (case e of
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   405
          q1 => ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   406
        | q2 => ...)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   407
    | foo p2 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   408
        let
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   409
          ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   410
        in
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   411
          ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   412
        end
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   413
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   414
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   415
  (* WRONG *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   416
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   417
  fun foo p1 = case e of
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   418
      q1 => ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   419
    | q2 => ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   420
    | foo p2 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   421
    let
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   422
      ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   423
    in
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   424
      ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   425
    end
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   426
  \end{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   427
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   428
  Extra parentheses around @{verbatim case} expressions are optional,
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   429
  but help to analyse the nesting easily based on simple string
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   430
  matching in the editor.
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   431
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   432
  \medskip There are two main exceptions to the overall principle of
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   433
  compositionality in the layout of complex expressions.
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   434
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   435
  \begin{enumerate}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   436
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   437
  \item @{verbatim "if"} expressions are iterated as if there would be
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   438
  a multi-branch conditional in SML, e.g.
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   439
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   440
  \begin{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   441
  (* RIGHT *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   442
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   443
  if b1 then e1
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   444
  else if b2 then e2
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   445
  else e3
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   446
  \end{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   447
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   448
  \item @{verbatim fn} abstractions are often layed-out as if they
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   449
  would lack any structure by themselves.  This traditional form is
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   450
  motivated by the possibility to shift function arguments back and
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   451
  forth wrt.\ additional combinators.  For example:
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   452
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   453
  \begin{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   454
  (* RIGHT *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   455
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   456
  fun foo x y = fold (fn z =>
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   457
    expr)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   458
  \end{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   459
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   460
  Here the visual appearance is that of three arguments @{verbatim x},
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   461
  @{verbatim y}, @{verbatim z}.
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   462
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   463
  \end{enumerate}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   464
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   465
  Such weakly structured layout should be use with great care.  Here
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   466
  is a counter-example involving @{verbatim let} expressions:
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   467
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   468
  \begin{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   469
  (* WRONG *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   470
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   471
  fun foo x = let
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   472
      val y = ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   473
    in ... end
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   474
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   475
  fun foo x =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   476
  let
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   477
    val y = ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   478
  in ... end
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   479
  \end{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   480
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   481
  \medskip In general the source layout is meant to emphasize the
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   482
  structure of complex language expressions, not to pretend that SML
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   483
  had a completely different syntax (say that of Haskell or Java).
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   484
*}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   485
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   486
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   487
section {* SML embedded into Isabelle/Isar *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   488
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   489
text {* ML and Isar are intertwined via an open-ended bootstrap
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   490
  process that provides more and more programming facilities and
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   491
  logical content in an alternating manner.  Bootstrapping starts from
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   492
  the raw environment of existing implementations of Standard ML
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   493
  (mainly Poly/ML, but also SML/NJ).
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   494
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   495
  Isabelle/Pure marks the point where the original ML toplevel is
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   496
  superseded by the Isar toplevel that maintains a uniform environment
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   497
  for arbitrary ML values (see also \secref{sec:context}).  This
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   498
  formal context holds logical entities as well as ML compiler
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   499
  bindings, among many other things.  Raw Standard ML is never
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   500
  encountered again after the initial bootstrap of Isabelle/Pure.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   501
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   502
  Object-logics such as Isabelle/HOL are built within the
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   503
  Isabelle/ML/Isar environment of Pure by introducing suitable
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   504
  theories with associated ML text, either inlined or as separate
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   505
  files.  Thus Isabelle/HOL is defined as a regular user-space
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   506
  application within the Isabelle framework.  Further add-on tools can
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   507
  be implemented in ML within the Isar context in the same manner: ML
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   508
  is part of the regular user-space repertoire of Isabelle.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   509
*}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   510
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   511
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   512
subsection {* Isar ML commands *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   513
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   514
text {* The primary Isar source language provides various facilities
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   515
  to open a ``window'' to the underlying ML compiler.  Especially see
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   516
  @{command_ref "use"} and @{command_ref "ML"}, which work exactly the
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   517
  same way, only the source text is provided via a file vs.\ inlined,
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   518
  respectively.  Apart from embedding ML into the main theory
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   519
  definition like that, there are many more commands that refer to ML
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   520
  source, such as @{command_ref setup} or @{command_ref declaration}.
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   521
  An example of even more fine-grained embedding of ML into Isar is
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   522
  the proof method @{method_ref tactic}, which refines the pending goal state
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   523
  via a given expression of type @{ML_type tactic}.
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   524
*}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   525
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   526
text %mlex {* The following artificial example demonstrates some ML
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   527
  toplevel declarations within the implicit Isar theory context.  This
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   528
  is regular functional programming without referring to logical
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   529
  entities yet.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   530
*}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   531
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   532
ML {*
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   533
  fun factorial 0 = 1
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   534
    | factorial n = n * factorial (n - 1)
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   535
*}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   536
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   537
text {* Here the ML environment is really managed by Isabelle, i.e.\
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   538
  the @{ML factorial} function is not yet accessible in the preceding
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   539
  paragraph, nor in a different theory that is independent from the
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   540
  current one in the import hierarchy.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   541
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   542
  Removing the above ML declaration from the source text will remove
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   543
  any trace of this definition as expected.  The Isabelle/ML toplevel
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   544
  environment is managed in a \emph{stateless} way: unlike the raw ML
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   545
  toplevel or similar command loops of Computer Algebra systems, there
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   546
  are no global side-effects involved here.\footnote{Such a stateless
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   547
  compilation environment is also a prerequisite for robust parallel
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   548
  compilation within independent nodes of the implicit theory
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   549
  development graph.}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   550
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   551
  \medskip The next example shows how to embed ML into Isar proofs.
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   552
  Instead of @{command_ref "ML"} for theory mode, we use @{command_ref
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   553
  "ML_prf"} for proof mode.  As illustrated below, its effect on the
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   554
  ML environment is local to the whole proof body, while ignoring its
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   555
  internal block structure.  *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   556
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   557
example_proof
39851
7219a771ab63 more examples;
wenzelm
parents: 39850
diff changeset
   558
  ML_prf %"ML" {* val a = 1 *}
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   559
  { -- {* Isar block structure ignored by ML environment *}
39851
7219a771ab63 more examples;
wenzelm
parents: 39850
diff changeset
   560
    ML_prf %"ML" {* val b = a + 1 *}
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   561
  } -- {* Isar block structure ignored by ML environment *}
39851
7219a771ab63 more examples;
wenzelm
parents: 39850
diff changeset
   562
  ML_prf %"ML" {* val c = b + 1 *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   563
qed
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   564
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   565
text {* By side-stepping the normal scoping rules for Isar proof
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   566
  blocks, embedded ML code can refer to the different contexts
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   567
  explicitly, and manipulate corresponding entities, e.g.\ export a
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   568
  fact from a block context.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   569
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   570
  \medskip Two further ML commands are useful in certain situations:
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   571
  @{command_ref ML_val} and @{command_ref ML_command} are both
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   572
  \emph{diagnostic} in the sense that there is no effect on the
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   573
  underlying environment, and can thus used anywhere (even outside a
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   574
  theory).  The examples below produce long strings of digits by
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   575
  invoking @{ML factorial}: @{command ML_val} already takes care of
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   576
  printing the ML toplevel result, but @{command ML_command} is silent
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   577
  so we produce an explicit output message.  *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   578
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   579
ML_val {* factorial 100 *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   580
ML_command {* writeln (string_of_int (factorial 100)) *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   581
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   582
example_proof
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   583
  ML_val {* factorial 100 *}  (* FIXME check/fix indentation *)
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   584
  ML_command {* writeln (string_of_int (factorial 100)) *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   585
qed
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   586
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   587
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   588
subsection {* Compile-time context *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   589
39825
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   590
text {* Whenever the ML compiler is invoked within Isabelle/Isar, the
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   591
  formal context is passed as a thread-local reference variable.  Thus
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   592
  ML code may access the theory context during compilation, by reading
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   593
  or writing the (local) theory under construction.  Note that such
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   594
  direct access to the compile-time context is rare; in practice it is
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   595
  typically via some derived ML functions.
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   596
*}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   597
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   598
text %mlref {*
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   599
  \begin{mldecls}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   600
  @{index_ML ML_Context.the_generic_context: "unit -> Context.generic"} \\
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   601
  @{index_ML "Context.>> ": "(Context.generic -> Context.generic) -> unit"} \\
39850
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   602
  @{index_ML bind_thms: "string * thm list -> unit"} \\
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   603
  @{index_ML bind_thm: "string * thm -> unit"} \\
39825
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   604
  \end{mldecls}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   605
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   606
  \begin{description}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   607
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   608
  \item @{ML "ML_Context.the_generic_context ()"} refers to the theory
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   609
  context of the ML toplevel --- at compile time.  ML code needs to
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   610
  take care to refer to @{ML "ML_Context.the_generic_context ()"}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   611
  correctly.  Recall that evaluation of a function body is delayed
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   612
  until actual run-time.
39825
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   613
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   614
  \item @{ML "Context.>>"}~@{text f} applies context transformation
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   615
  @{text f} to the implicit context of the ML toplevel.
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   616
39850
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   617
  \item @{ML bind_thms}~@{text "(name, thms)"} stores a list of
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   618
  theorems produced in ML both in the (global) theory context and the
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   619
  ML toplevel, associating it with the provided name.  Theorems are
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   620
  put into a global ``standard'' format before being stored.
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   621
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   622
  \item @{ML bind_thm} is similar to @{ML bind_thms} but refers to a
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   623
  singleton theorem.
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   624
39825
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   625
  \end{description}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   626
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   627
  It is very important to note that the above functions are really
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   628
  restricted to the compile time, even though the ML compiler is
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   629
  invoked at run-time.  The majority of ML code either uses static
39825
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   630
  antiquotations (\secref{sec:ML-antiq}) or refers to the theory or
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   631
  proof context at run-time, by explicit functional abstraction.
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   632
*}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   633
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   634
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   635
subsection {* Antiquotations \label{sec:ML-antiq} *}
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   636
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   637
text {* A very important consequence of embedding SML into Isar is the
39829
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   638
  concept of \emph{ML antiquotation}.  First, the standard token
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   639
  language of ML is augmented by special syntactic entities of the
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   640
  following form:
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   641
39829
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   642
  \indexouternonterm{antiquote}
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   643
  \begin{rail}
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   644
  antiquote: atsign lbrace nameref args rbrace | lbracesym | rbracesym
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   645
  ;
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   646
  \end{rail}
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   647
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   648
  Here @{syntax nameref} and @{syntax args} are regular outer syntax
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   649
  categories.  Note that attributes and proof methods use similar
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   650
  syntax.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   651
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   652
  \medskip A regular antiquotation @{text "@{name args}"} processes
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   653
  its arguments by the usual means of the Isar source language, and
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   654
  produces corresponding ML source text, either as literal
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   655
  \emph{inline} text (e.g. @{text "@{term t}"}) or abstract
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   656
  \emph{value} (e.g. @{text "@{thm th}"}).  This pre-compilation
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   657
  scheme allows to refer to formal entities in a robust manner, with
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   658
  proper static scoping and with some degree of logical checking of
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   659
  small portions of the code.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   660
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   661
  Special antiquotations like @{text "@{let \<dots>}"} or @{text "@{note
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   662
  \<dots>}"} augment the compilation context without generating code.  The
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   663
  non-ASCII braces @{text "\<lbrace>"} and @{text "\<rbrace>"} allow to delimit the
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   664
  effect by introducing local blocks within the pre-compilation
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   665
  environment.
39829
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   666
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   667
  \medskip See also \cite{Wenzel-Chaieb:2007b} for a slightly broader
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   668
  perspective on Isabelle/ML antiquotations.  *}
39829
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   669
39830
7c501d7f1e45 clarified tag markup;
wenzelm
parents: 39829
diff changeset
   670
text %mlantiq {*
39829
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   671
  \begin{matharray}{rcl}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   672
  @{ML_antiquotation_def "let"} & : & @{text ML_antiquotation} \\
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   673
  @{ML_antiquotation_def "note"} & : & @{text ML_antiquotation} \\
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   674
  \end{matharray}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   675
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   676
  \begin{rail}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   677
  'let' ((term + 'and') '=' term + 'and')
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   678
  ;
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   679
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   680
  'note' (thmdef? thmrefs + 'and')
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   681
  ;
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   682
  \end{rail}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   683
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   684
  \begin{description}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   685
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   686
  \item @{text "@{let p = t}"} binds schematic variables in the
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   687
  pattern @{text "p"} by higher-order matching against the term @{text
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   688
  "t"}.  This is analogous to the regular @{command_ref let} command
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   689
  in the Isar proof language.  The pre-compilation environment is
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   690
  augmented by auxiliary term bindings, without emitting ML source.
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   691
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   692
  \item @{text "@{note a = b\<^sub>1 \<dots> b\<^sub>n}"} recalls existing facts @{text
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   693
  "b\<^sub>1, \<dots>, b\<^sub>n"}, binding the result as @{text a}.  This is analogous to
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   694
  the regular @{command_ref note} command in the Isar proof language.
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   695
  The pre-compilation environment is augmented by auxiliary fact
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   696
  bindings, without emitting ML source.
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   697
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   698
  \end{description}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   699
*}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   700
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   701
text %mlex {* The following artificial example gives a first
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   702
  impression about using the antiquotation elements introduced so far,
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   703
  together with the basic @{text "@{thm}"} antiquotation defined
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   704
  later.
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   705
*}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   706
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   707
ML {*
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   708
  \<lbrace>
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   709
    @{let ?t = my_term}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   710
    @{note my_refl = reflexive [of ?t]}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   711
    fun foo th = Thm.transitive th @{thm my_refl}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   712
  \<rbrace>
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   713
*}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   714
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   715
text {*
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   716
  The extra block delimiters do not affect the compiled code itself, i.e.\
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   717
  function @{ML foo} is available in the present context.
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   718
*}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   719
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   720
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   721
section {* Message output channels \label{sec:message-channels} *}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   722
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   723
text {* Isabelle provides output channels for different kinds of
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   724
  messages: regular output, high-volume tracing information, warnings,
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   725
  and errors.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   726
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   727
  Depending on the user interface involved, these messages may appear
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   728
  in different text styles or colours.  The standard output for
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   729
  terminal sessions prefixes each line of warnings by @{verbatim
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   730
  "###"} and errors by @{verbatim "***"}, but leaves anything else
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   731
  unchanged.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   732
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   733
  Messages are associated with the transaction context of the running
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   734
  Isar command.  This enables the front-end to manage commands and
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   735
  resulting messages together.  For example, after deleting a command
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   736
  from a given theory document version, the corresponding message
39872
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   737
  output can be retracted from the display.
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   738
*}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   739
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   740
text %mlref {*
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   741
  \begin{mldecls}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   742
  @{index_ML writeln: "string -> unit"} \\
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   743
  @{index_ML tracing: "string -> unit"} \\
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   744
  @{index_ML warning: "string -> unit"} \\
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   745
  @{index_ML error: "string -> 'a"} \\
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   746
  \end{mldecls}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   747
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   748
  \begin{description}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   749
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   750
  \item @{ML writeln}~@{text "text"} outputs @{text "text"} as regular
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   751
  message.  This is the primary message output operation of Isabelle
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   752
  and should be used by default.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   753
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   754
  \item @{ML tracing}~@{text "text"} outputs @{text "text"} as special
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   755
  tracing message, indicating potential high-volume output to the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   756
  front-end (hundreds or thousands of messages issued by a single
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   757
  command).  The idea is to allow the user-interface to downgrade the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   758
  quality of message display to achieve higher throughput.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   759
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   760
  Note that the user might have to take special actions to see tracing
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   761
  output, e.g.\ switch to a different output window.  So this channel
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   762
  should not be used for regular output.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   763
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   764
  \item @{ML warning}~@{text "text"} outputs @{text "text"} as
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   765
  warning, which typically means some extra emphasis on the front-end
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   766
  side (color highlighting, icon).
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   767
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   768
  \item @{ML error}~@{text "text"} raises exception @{ML ERROR}~@{text
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   769
  "text"} and thus lets the Isar toplevel print @{text "text"} on the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   770
  error channel, which typically means some extra emphasis on the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   771
  front-end side (color highlighting, icon).
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   772
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   773
  This assumes that the exception is not handled before the command
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   774
  terminates.  Handling exception @{ML ERROR}~@{text "text"} is a
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   775
  perfectly legal alternative: it means that the error is absorbed
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   776
  without any message output.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   777
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   778
  \begin{warn}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   779
  The actual error channel is accessed via @{ML Output.error_msg}, but
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   780
  the interaction protocol of Proof~General \emph{crashes} if that
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   781
  function is used in regular ML code: error output and toplevel
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   782
  command failure always need to coincide here.
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   783
  \end{warn}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   784
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   785
  \end{description}
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   786
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   787
  \begin{warn}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   788
  Regular Isabelle/ML code should output messages exclusively by the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   789
  official channels.  Using raw I/O on \emph{stdout} or \emph{stderr}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   790
  instead (e.g.\ via @{ML TextIO.output}) is apt to cause problems in
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   791
  the presence of parallel and asynchronous processing of Isabelle
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   792
  theories.  Such raw output might be displayed by the front-end in
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   793
  some system console log, with a low chance that the user will ever
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   794
  see it.  Moreover, as a genuine side-effect on global process
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   795
  channels, there is no proper way to retract output when Isar command
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   796
  transactions are reset.
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   797
  \end{warn}
39872
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   798
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   799
  \begin{warn}
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   800
  The message channels should be used in a message-oriented manner.
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   801
  This means that multi-line output that logically belongs together
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   802
  needs to be issued by a \emph{single} invocation of @{ML writeln}
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   803
  etc.  with the functional concatenation of all message constituents.
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   804
  \end{warn}
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   805
*}
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   806
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   807
text %mlex {* The following example demonstrates a multi-line
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   808
  warning.  Note that in some situations the user sees only the first
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   809
  line, so the most important point should be made first.
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   810
*}
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   811
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   812
ML_command {*
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   813
  warning (cat_lines
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   814
   ["Beware the Jabberwock, my son!",
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   815
    "The jaws that bite, the claws that catch!",
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   816
    "Beware the Jubjub Bird, and shun",
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
   817
    "The frumious Bandersnatch!"]);
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   818
*}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   819
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   820
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   821
section {* Exceptions \label{sec:exceptions} *}
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   822
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   823
text {* The Standard ML semantics of strict functional evaluation
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   824
  together with exceptions is rather well defined, but some delicate
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   825
  points need to be observed to avoid that ML programs go wrong
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   826
  despite static type-checking.  Exceptions in Isabelle/ML are
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   827
  subsequently categorized as follows.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   828
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   829
  \paragraph{Regular user errors.}  These are meant to provide
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   830
  informative feedback about malformed input etc.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   831
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   832
  The \emph{error} function raises the corresponding \emph{ERROR}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   833
  exception, with a plain text message as argument.  \emph{ERROR}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   834
  exceptions can be handled internally, in order to be ignored, turned
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   835
  into other exceptions, or cascaded by appending messages.  If the
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   836
  corresponding Isabelle/Isar command terminates with an \emph{ERROR}
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   837
  exception state, the toplevel will print the result on the error
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   838
  channel (see \secref{sec:message-channels}).
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   839
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   840
  It is considered bad style to refer to internal function names or
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   841
  values in ML source notation in user error messages.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   842
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   843
  Grammatical correctness of error messages can be improved by
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   844
  \emph{omitting} final punctuation: messages are often concatenated
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   845
  or put into a larger context (e.g.\ augmented with source position).
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   846
  By not insisting in the final word at the origin of the error, the
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   847
  system can perform its administrative tasks more easily and
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   848
  robustly.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   849
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   850
  \paragraph{Program failures.}  There is a handful of standard
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   851
  exceptions that indicate general failure situations, or failures of
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   852
  core operations on logical entities (types, terms, theorems,
39856
wenzelm
parents: 39855
diff changeset
   853
  theories, see \chref{ch:logic}).
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   855
  These exceptions indicate a genuine breakdown of the program, so the
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   856
  main purpose is to determine quickly what has happened where.
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   857
  Traditionally, the (short) exception message would include the name
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   858
  of an ML function, although this no longer necessary, because the ML
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   859
  runtime system prints a detailed source position of the
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   860
  corresponding @{verbatim raise} keyword.
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   861
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   862
  \medskip User modules can always introduce their own custom
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   863
  exceptions locally, e.g.\ to organize internal failures robustly
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   864
  without overlapping with existing exceptions.  Exceptions that are
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   865
  exposed in module signatures require extra care, though, and should
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   866
  \emph{not} be introduced by default.  Surprise by end-users or ML
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   867
  users of a module can be often minimized by using plain user errors.
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   868
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   869
  \paragraph{Interrupts.}  These indicate arbitrary system events:
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   870
  both the ML runtime system and the Isabelle/ML infrastructure signal
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   871
  various exceptional situations by raising the special
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   872
  \emph{Interrupt} exception in user code.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   873
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   874
  This is the one and only way that physical events can intrude an
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   875
  Isabelle/ML program.  Such an interrupt can mean out-of-memory,
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   876
  stack overflow, timeout, internal signaling of threads, or the user
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   877
  producing a console interrupt manually etc.  An Isabelle/ML program
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   878
  that intercepts interrupts becomes dependent on physical effects of
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   879
  the environment.  Even worse, exception handling patterns that are
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   880
  too general by accident, e.g.\ by mispelled exception constructors,
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   881
  will cover interrupts unintentionally, and thus render the program
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   882
  semantics ill-defined.
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   883
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   884
  Note that the Interrupt exception dates back to the original SML90
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   885
  language definition.  It was excluded from the SML97 version to
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   886
  avoid its malign impact on ML program semantics, but without
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   887
  providing a viable alternative.  Isabelle/ML recovers physical
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   888
  interruptibility (which an indispensable tool to implement managed
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   889
  evaluation of Isar command transactions), but requires user code to
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   890
  be strictly transparent wrt.\ interrupts.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   891
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   892
  \begin{warn}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   893
  Isabelle/ML user code needs to terminate promptly on interruption,
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   894
  without guessing at its meaning to the system infrastructure.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   895
  Temporary handling of interrupts for cleanup of global resources
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   896
  etc.\ needs to be followed immediately by re-raising of the original
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   897
  exception.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   898
  \end{warn}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   899
*}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   900
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   901
text %mlref {*
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   902
  \begin{mldecls}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   903
  @{index_ML try: "('a -> 'b) -> 'a -> 'b option"} \\
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   904
  @{index_ML can: "('a -> 'b) -> 'a -> bool"} \\
39856
wenzelm
parents: 39855
diff changeset
   905
  @{index_ML ERROR: "string -> exn"} \\
wenzelm
parents: 39855
diff changeset
   906
  @{index_ML Fail: "string -> exn"} \\
wenzelm
parents: 39855
diff changeset
   907
  @{index_ML Exn.is_interrupt: "exn -> bool"} \\
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   908
  @{index_ML reraise: "exn -> 'a"} \\
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   909
  @{index_ML exception_trace: "(unit -> 'a) -> 'a"} \\
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   910
  \end{mldecls}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   911
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   912
  \begin{description}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   913
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   914
  \item @{ML try}~@{text "f x"} makes the partiality of evaluating
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   915
  @{text "f x"} explicit via the option datatype.  Interrupts are
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   916
  \emph{not} handled here, i.e.\ this form serves as safe replacement
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   917
  for the \emph{unsafe} version @{verbatim "(SOME"}~@{text "f
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   918
  x"}~@{verbatim "handle _ => NONE)"} that is occasionally seen in
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   919
  books about SML.
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   920
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   921
  \item @{ML can} is similar to @{ML try} with more abstract result.
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   922
39856
wenzelm
parents: 39855
diff changeset
   923
  \item @{ML ERROR}~@{text "msg"} represents user errors; this
wenzelm
parents: 39855
diff changeset
   924
  exception is always raised via the @{ML error} function (see
wenzelm
parents: 39855
diff changeset
   925
  \secref{sec:message-channels}).
wenzelm
parents: 39855
diff changeset
   926
wenzelm
parents: 39855
diff changeset
   927
  \item @{ML Fail}~@{text "msg"} represents general program failures.
wenzelm
parents: 39855
diff changeset
   928
wenzelm
parents: 39855
diff changeset
   929
  \item @{ML Exn.is_interrupt} identifies interrupts robustly, without
wenzelm
parents: 39855
diff changeset
   930
  mentioning concrete exception constructors in user code.  Handled
wenzelm
parents: 39855
diff changeset
   931
  interrupts need to be re-raised promptly!
wenzelm
parents: 39855
diff changeset
   932
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   933
  \item @{ML reraise}~@{text "exn"} raises exception @{text "exn"}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   934
  while preserving its implicit position information (if possible,
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   935
  depending on the ML platform).
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   936
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   937
  \item @{ML exception_trace}~@{verbatim "(fn _ =>"}~@{text
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   938
  "e"}@{verbatim ")"} evaluates expression @{text "e"} while printing
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   939
  a full trace of its stack of nested exceptions (if possible,
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   940
  depending on the ML platform).\footnote{In various versions of
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   941
  Poly/ML the trace will appear on raw stdout of the Isabelle
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   942
  process.}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   943
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   944
  Inserting @{ML exception_trace} into ML code temporarily is useful
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   945
  for debugging, but not suitable for production code.
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   946
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   947
  \end{description}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   948
*}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   949
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   950
text %mlantiq {*
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   951
  \begin{matharray}{rcl}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   952
  @{ML_antiquotation_def "assert"} & : & @{text ML_antiquotation} \\
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   953
  \end{matharray}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   954
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   955
  \begin{description}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   956
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   957
  \item @{text "@{assert}"} inlines a function @{text "bool \<Rightarrow> unit"}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   958
  that raises @{ML Fail} if the argument is @{ML false}.  Due to
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   959
  inlining the source position of failed assertions is included in the
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   960
  error output.
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   961
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   962
  \end{description}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   963
*}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   964
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   965
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   966
section {* Basic data types *}
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   967
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   968
text {* The basis library proposal of SML97 need to be treated with
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   969
  caution.  Many of its operations simply do not fit with important
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   970
  Isabelle/ML conventions (like ``canonical argument order'', see
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   971
  \secref{sec:canonical-argument-order}), others can cause serious
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   972
  problems with the parallel evaluation model of Isabelle/ML (such as
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   973
  @{ML TextIO.print} or @{ML OS.Process.system}).
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   974
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   975
  Subsequently we give a brief overview of important operations on
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   976
  basic ML data types.
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   977
*}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   978
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   979
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   980
subsection {* Characters *}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   981
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   982
text %mlref {*
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   983
  \begin{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   984
  @{index_ML_type char} \\
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   985
  \end{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   986
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   987
  \begin{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   988
39864
wenzelm
parents: 39863
diff changeset
   989
  \item Type @{ML_type char} is \emph{not} used.  The smallest textual
wenzelm
parents: 39863
diff changeset
   990
  unit in Isabelle is represented a ``symbol'' (see
wenzelm
parents: 39863
diff changeset
   991
  \secref{sec:symbols}).
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   992
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   993
  \end{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   994
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   995
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   996
39862
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   997
subsection {* Integers *}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   998
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   999
text %mlref {*
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1000
  \begin{mldecls}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1001
  @{index_ML_type int} \\
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1002
  \end{mldecls}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1003
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1004
  \begin{description}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1005
39864
wenzelm
parents: 39863
diff changeset
  1006
  \item Type @{ML_type int} represents regular mathematical integers,
wenzelm
parents: 39863
diff changeset
  1007
  which are \emph{unbounded}.  Overflow never happens in
39862
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1008
  practice.\footnote{The size limit for integer bit patterns in memory
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1009
  is 64\,MB for 32-bit Poly/ML, and much higher for 64-bit systems.}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1010
  This works uniformly for all supported ML platforms (Poly/ML and
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1011
  SML/NJ).
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1012
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1013
  Literal integers in ML text (e.g.\ @{ML
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1014
  123456789012345678901234567890}) are forced to be of this one true
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1015
  integer type --- overloading of SML97 is disabled.
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1016
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1017
  Structure @{ML_struct IntInf} of SML97 is obsolete, always use
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1018
  @{ML_struct Int}.  Structure @{ML_struct Integer} in @{"file"
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1019
  "~~/src/Pure/General/integer.ML"} provides some additional
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1020
  operations.
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1021
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1022
  \end{description}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1023
*}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1024
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1025
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1026
subsection {* Options *}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1027
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1028
text %mlref {*
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1029
  \begin{mldecls}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1030
  @{index_ML Option.map: "('a -> 'b) -> 'a option -> 'b option"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1031
  @{index_ML is_some: "'a option -> bool"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1032
  @{index_ML is_none: "'a option -> bool"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1033
  @{index_ML the: "'a option -> 'a"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1034
  @{index_ML these: "'a list option -> 'a list"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1035
  @{index_ML the_list: "'a option -> 'a list"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1036
  @{index_ML the_default: "'a -> 'a option -> 'a"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1037
  \end{mldecls}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1038
*}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1039
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1040
text {* Apart from @{ML Option.map} most operations defined in
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1041
  structure @{ML_struct Option} are alien to Isabelle/ML.  The
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1042
  operations shown above are defined in @{"file"
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1043
  "~~/src/Pure/General/basics.ML"}, among others.  *}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1044
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1045
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1046
subsection {* Lists *}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1047
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1048
text {* Lists are ubiquitous in ML as simple and light-weight
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1049
  ``collections'' for many everyday programming tasks.  Isabelle/ML
39874
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1050
  provides important additions and improvements over operations that
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1051
  are predefined in the SML97 library. *}
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1052
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1053
text %mlref {*
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1054
  \begin{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1055
  @{index_ML cons: "'a -> 'a list -> 'a list"} \\
39874
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1056
  @{index_ML member: "('b * 'a -> bool) -> 'a list -> 'b -> bool"} \\
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1057
  @{index_ML insert: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1058
  @{index_ML remove: "('b * 'a -> bool) -> 'b -> 'a list -> 'a list"} \\
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1059
  @{index_ML update: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1060
  \end{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1061
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1062
  \begin{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1063
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1064
  \item @{ML cons}~@{text "x xs"} evaluates to @{text "x :: xs"}.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1065
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1066
  Tupled infix operators are a historical accident in Standard ML.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1067
  The curried @{ML cons} amends this, but it should be only used when
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1068
  partial application is required.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1069
39874
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1070
  \item @{ML member}, @{ML insert}, @{ML remove}, @{ML update} treat
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1071
  lists as a set-like container that maintains the order of elements.
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1072
  See @{"file" "~~/src/Pure/library.ML"} for the full specifications
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1073
  (written in ML).  There are some further derived operations like
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1074
  @{ML union} or @{ML inter}.
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1075
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1076
  Note that @{ML insert} is conservative about elements that are
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1077
  already a @{ML member} of the list, while @{ML update} ensures that
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1078
  the last entry is always put in front.  The latter discipline is
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1079
  often more appropriate in declarations of context data
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1080
  (\secref{sec:context-data}) that are issued by the user in Isar
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1081
  source: more recent declarations normally take precedence over
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1082
  earlier ones.
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1083
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1084
  \end{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1085
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1086
39874
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1087
text %mlex {* The following example demonstrates how to \emph{merge}
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1088
  two lists in a natural way. *}
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1089
39874
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1090
ML {*
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1091
  fun merge_lists eq (xs, ys) = fold_rev (insert eq) ys xs;
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1092
*}
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1093
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1094
text {* Here the first list is treated conservatively: only the new
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1095
  elements from the second list are inserted.  The inside-out order of
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1096
  insertion via @{ML fold_rev} attempts to preserve the order of
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1097
  elements in the result.
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1098
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1099
  This way of merging lists is typical for context data
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1100
  (\secref{sec:context-data}).  See also @{ML merge} as defined in
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1101
  @{"file" "~~/src/Pure/library.ML"}.
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1102
*}
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1103
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1104
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1105
subsubsection {* Canonical iteration *}  (* FIXME move!? *)
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1106
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1107
text {* A function @{text "f: \<alpha> \<rightarrow> \<beta> \<rightarrow> \<beta>"} can be understood as update
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1108
  on a configuration of type @{text "\<beta>"} that is parametrized by
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1109
  arguments of type @{text "\<alpha>"}.  Given @{text "a: \<alpha>"} the partial
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1110
  application @{text "(f a): \<beta> \<rightarrow> \<beta>"} operates homogeneously on @{text
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1111
  "\<beta>"}.  This can be iterated naturally over a list of parameters
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1112
  @{text "[a\<^sub>1, \<dots>, a\<^sub>n]"} as @{text "f a\<^sub>1; \<dots>; f a\<^bsub>n\<^esub>\<^bsub>\<^esub>"} (where the
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1113
  semicolon represents left-to-right composition).  The latter
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1114
  expression is again a function @{text "\<beta> \<rightarrow> \<beta>"}.  It can be applied
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1115
  to an initial configuration @{text "b: \<beta>"} to start the iteration
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1116
  over the given list of arguments: each @{text "a"} in @{text "a\<^sub>1, \<dots>,
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1117
  a\<^sub>n"} is applied consecutively by updating a cumulative
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1118
  configuration.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1119
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1120
  The @{text fold} combinator in Isabelle/ML lifts a function @{text
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1121
  "f"} as above to its iterated version over a list of arguments.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1122
  Lifting can be repeated, e.g.\ @{text "(fold \<circ> fold) f"} iterates
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1123
  over a list of lists as expected.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1124
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1125
  The variant @{text "fold_rev"} works inside-out over the list of
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1126
  arguments, such that @{text "fold_rev f \<equiv> fold f \<circ> rev"} holds.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1127
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1128
  The @{text "fold_map"} combinator essentially performs @{text
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1129
  "fold"} and @{text "map"} simultaneously: each application of @{text
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1130
  "f"} produces an updated configuration together with a side-result;
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1131
  the iteration collects all such side-results as a separate list.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1132
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1133
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1134
text %mlref {*
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1135
  \begin{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1136
  @{index_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1137
  @{index_ML fold_rev: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1138
  @{index_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1139
  \end{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1140
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1141
  \begin{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1142
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1143
  \item @{ML fold}~@{text f} lifts the parametrized update function
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1144
  @{text "f"} to a list of parameters.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1145
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1146
  \item @{ML fold_rev}~@{text "f"} is similar to @{ML fold}~@{text
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1147
  "f"}, but works inside-out.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1148
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1149
  \item @{ML fold_map}~@{text "f"} lifts the parametrized update
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1150
  function @{text "f"} (with side-result) to a list of parameters and
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1151
  cumulative side-results.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1152
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1153
  \end{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1154
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1155
  \begin{warn}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1156
  The literature on functional programming provides a multitude of
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1157
  combinators called @{text "foldl"}, @{text "foldr"} etc.  SML97
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1158
  provides its own variations as @{ML List.foldl} and @{ML
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1159
  List.foldr}, while the classic Isabelle library still has the
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1160
  slightly more convenient historic @{ML Library.foldl} and @{ML
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1161
  Library.foldr}.  To avoid further confusion, all of this should be
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1162
  ignored, and @{ML fold} (or @{ML fold_rev}) used exclusively.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1163
  \end{warn}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1164
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1165
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1166
text %mlex {* Using canonical @{ML fold} together with canonical @{ML
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1167
  cons}, or similar standard operations, alternates the orientation of
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1168
  data.  The is quite natural and should not altered forcible by
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1169
  inserting extra applications @{ML rev}.  The alternative @{ML
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1170
  fold_rev} can be used in the relatively few situations, where
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1171
  alternation should be prevented.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1172
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1173
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1174
ML {*
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1175
  val items = [1, 2, 3, 4, 5, 6, 7, 8, 9, 10];
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1176
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1177
  val list1 = fold cons items [];
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1178
  @{assert} (list1 = rev items);
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1179
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1180
  val list2 = fold_rev cons items [];
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1181
  @{assert} (list2 = items);
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1182
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1183
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1184
39875
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1185
subsection {* Association lists *}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1186
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1187
text {* The operations for association lists interpret a concrete list
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1188
  of pairs as a finite function from keys to values.  Redundant
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1189
  representations with multiple occurrences of the same key are
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1190
  implicitly normalized: lookup and update only take the first
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1191
  occurrence into account.
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1192
*}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1193
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1194
text {*
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1195
  \begin{mldecls}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1196
  @{index_ML AList.lookup: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option"} \\
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1197
  @{index_ML AList.defined: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool"} \\
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1198
  @{index_ML AList.update: "('a * 'a -> bool) -> 'a * 'b -> ('a * 'b) list -> ('a * 'b) list"} \\
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1199
  \end{mldecls}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1200
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1201
  \begin{description}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1202
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1203
  \item @{ML AList.lookup}, @{ML AList.defined}, @{ML AList.update}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1204
  implement the main ``framework operations'' for mappings in
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1205
  Isabelle/ML, with standard conventions for names and types.
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1206
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1207
  Note that a function called @{text lookup} is obliged to express its
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1208
  partiality via an explicit option element.  There is no choice to
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1209
  raise an exception, without changing the name to something like
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1210
  @{text "the_element"} or @{text "get"}.
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1211
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1212
  The @{text "defined"} operation is essentially a contraction of @{ML
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1213
  is_some} and @{text "lookup"}, but this is sufficiently frequent to
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1214
  justify its independent existence.  This also gives the
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1215
  implementation some opportunity for peep-hole optimization.
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1216
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1217
  \end{description}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1218
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1219
  Association lists are adequate as simple and light-weight
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1220
  implementation of finite mappings in many practical situations.  A
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1221
  more heavy-duty table structure is defined in @{"file"
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1222
  "~~/src/Pure/General/table.ML"}; that version scales easily to
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1223
  thousands or millions of elements.
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1224
*}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1225
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1226
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1227
subsection {* Unsynchronized references *}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1228
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1229
text %mlref {*
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1230
  \begin{mldecls}
39870
wenzelm
parents: 39868
diff changeset
  1231
  @{index_ML_type "'a Unsynchronized.ref"} \\
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1232
  @{index_ML Unsynchronized.ref: "'a -> 'a Unsynchronized.ref"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1233
  @{index_ML "!": "'a Unsynchronized.ref -> 'a"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1234
  @{index_ML "op :=": "'a Unsynchronized.ref * 'a -> unit"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1235
  \end{mldecls}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1236
*}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1237
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1238
text {* Due to ubiquitous parallelism in Isabelle/ML (see also
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1239
  \secref{sec:multi-threading}), the mutable reference cells of
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1240
  Standard ML are notorious for causing problems.  In a highly
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1241
  parallel system, both correctness \emph{and} performance are easily
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1242
  degraded when using mutable data.
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1243
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1244
  The unwieldy name of @{ML Unsynchronized.ref} for the constructor
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1245
  for references in Isabelle/ML emphasizes the inconveniences caused by
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1246
  mutability.  Existing operations @{ML "!"}  and @{ML "op :="} are
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1247
  unchanged, but should be used with special precautions, say in a
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1248
  strictly local situation that is guaranteed to be restricted to
39870
wenzelm
parents: 39868
diff changeset
  1249
  sequential evaluation --- now and in the future.  *}
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1250
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1251
39870
wenzelm
parents: 39868
diff changeset
  1252
section {* Thread-safe programming \label{sec:multi-threading} *}
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1253
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1254
text {* Multi-threaded execution has become an everyday reality in
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1255
  Isabelle since Poly/ML 5.2.1 and Isabelle2008.  Isabelle/ML provides
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1256
  implicit and explicit parallelism by default, and there is no way
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1257
  for user-space tools to ``opt out''.  ML programs that are purely
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1258
  functional, output messages only via the official channels
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1259
  (\secref{sec:message-channels}), and do not intercept interrupts
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1260
  (\secref{sec:exceptions}) can participate in the multi-threaded
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1261
  environment immediately without further ado.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1262
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1263
  More ambitious tools with more fine-grained interaction with the
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1264
  environment need to observe the principles explained below.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1265
*}
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1266
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1267
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1268
subsection {* Multi-threading with shared memory *}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1269
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1270
text {* Multiple threads help to organize advanced operations of the
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1271
  system, such as real-time conditions on command transactions,
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1272
  sub-components with explicit communication, general asynchronous
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1273
  interaction etc.  Moreover, parallel evaluation is a prerequisite to
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1274
  make adequate use of the CPU resources that are available on
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1275
  multi-core systems.\footnote{Multi-core computing does not mean that
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1276
  there are ``spare cycles'' to be wasted.  It means that the
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1277
  continued exponential speedup of CPU performance due to ``Moore's
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1278
  Law'' follows different rules: clock frequency has reached its peak
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1279
  around 2005, and applications need to be parallelized in order to
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1280
  avoid a perceived loss of performance.  See also
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1281
  \cite{Sutter:2005}.}
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1282
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1283
  Isabelle/Isar exploits the inherent structure of theories and proofs
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1284
  to support \emph{implicit parallelism} to a large extent.  LCF-style
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1285
  theorem provides almost ideal conditions for that; see also
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1286
  \cite{Wenzel:2009}.  This means, significant parts of theory and
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1287
  proof checking is parallelized by default.  A maximum speedup-factor
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1288
  of 3.0 on 4 cores and 5.0 on 8 cores can be
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1289
  expected.\footnote{Further scalability is limited due to garbage
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1290
  collection, which is still sequential in Poly/ML 5.2/5.3/5.4.  It
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1291
  helps to provide initial heap space generously, using the
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1292
  \texttt{-H} option.  Initial heap size needs to be scaled-up
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1293
  together with the number of CPU cores: approximately 1--2\,GB per
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1294
  core..}
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1295
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1296
  \medskip ML threads lack the memory protection of separate
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1297
  processes, and operate concurrently on shared heap memory.  This has
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1298
  the advantage that results of independent computations are
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1299
  immediately available to other threads: abstract values can be
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1300
  passed between threads without copying or awkward serialization that
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1301
  is typically required for explicit message passing between separate
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1302
  processes.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1303
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1304
  To make shared-memory multi-threading work robustly and efficiently,
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1305
  some programming guidelines need to be observed.  While the ML
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1306
  system is responsible to maintain basic integrity of the
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1307
  representation of ML values in memory, the application programmer
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1308
  needs to ensure that multi-threaded execution does not break the
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1309
  intended semantics.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1310
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1311
  \begin{warn}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1312
  To participate in implicit parallelism, tools need to be
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1313
  thread-safe.  A single ill-behaved tool can affect the stability and
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1314
  performance of the whole system.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1315
  \end{warn}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1316
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1317
  Apart from observing the principles of thread-safeness passively,
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1318
  advanced tools may also exploit parallelism actively, e.g.\ by using
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1319
  ``future values'' (\secref{sec:futures}) or the more basic library
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1320
  functions for parallel list operations (\secref{sec:parlist}).
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1321
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1322
  \begin{warn}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1323
  Parallel computing resources are managed centrally by the
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1324
  Isabelle/ML infrastructure.  User programs must not fork their own
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1325
  ML threads to perform computations.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1326
  \end{warn}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1327
*}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1328
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1329
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1330
subsection {* Critical shared resources *}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1331
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1332
text {* Thread-safeness is mainly concerned about concurrent
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1333
  read/write access to shared resources, which are outside the purely
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1334
  functional world of ML.  This covers the following in particular.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1335
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1336
  \begin{itemize}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1337
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1338
  \item Global references (or arrays), i.e.\ mutable memory cells that
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1339
  persist over several invocations of associated
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1340
  operations.\footnote{This is independent of the visibility of such
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1341
  mutable values in the toplevel scope.}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1342
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1343
  \item Global state of the running Isabelle/ML process, i.e.\ raw I/O
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1344
  channels, environment variables, current working directory.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1345
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1346
  \item Writable resources in the file-system that are shared among
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1347
  different threads or other processes.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1348
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1349
  \end{itemize}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1350
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1351
  Isabelle/ML provides various mechanisms to avoid critical shared
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1352
  resources in most practical situations.  As last resort there are
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1353
  some mechanisms for explicit synchronization.  The following
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1354
  guidelines help to make Isabelle/ML programs work smoothly in a
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1355
  concurrent environment.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1356
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1357
  \begin{itemize}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1358
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1359
  \item Avoid global references altogether.  Isabelle/Isar maintains a
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1360
  uniform context that incorporates arbitrary data declared by user
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1361
  programs (\secref{sec:context-data}).  This context is passed as
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1362
  plain value and user tools can get/map their own data in a purely
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1363
  functional manner.  Configuration options within the context
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1364
  (\secref{sec:config-options}) provide simple drop-in replacements
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1365
  for formerly writable flags.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1366
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1367
  \item Keep components with local state information re-entrant.
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1368
  Instead of poking initial values into (private) global references, a
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1369
  new state record can be created on each invocation, and passed
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1370
  through any auxiliary functions of the component.  The state record
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1371
  may well contain mutable references, without requiring any special
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1372
  synchronizations, as long as each invocation gets its own copy.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1373
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1374
  \item Avoid raw output on @{text "stdout"} or @{text "stderr"}.  The
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1375
  Poly/ML library is thread-safe for each individual output operation,
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1376
  but the ordering of parallel invocations is arbitrary.  This means
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1377
  raw output will appear on some system console with unpredictable
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1378
  interleaving of atomic chunks.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1379
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1380
  Note that this does not affect regular message output channels
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1381
  (\secref{sec:message-channels}).  An official message is associated
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1382
  with the command transaction from where it originates, independently
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1383
  of other transactions.  This means each running Isar command has
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1384
  effectively its own set of message channels, and interleaving can
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1385
  only happen when commands use parallelism internally (and only at
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1386
  message boundaries).
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1387
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1388
  \item Treat environment variables and the current working directory
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1389
  of the running process as strictly read-only.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1390
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1391
  \item Restrict writing to the file-system to unique temporary files.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1392
  Isabelle already provides a temporary directory that is unique for
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1393
  the running process, and there is a centralized source of unique
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1394
  serial numbers in Isabelle/ML.  Thus temporary files that are passed
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1395
  to to some external process will be always disjoint, and thus
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1396
  thread-safe.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1397
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1398
  \end{itemize}
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1399
*}
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1400
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1401
text %mlref {*
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1402
  \begin{mldecls}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1403
  @{index_ML File.tmp_path: "Path.T -> Path.T"} \\
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1404
  @{index_ML serial_string: "unit -> string"} \\
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1405
  \end{mldecls}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1406
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1407
  \begin{description}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1408
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1409
  \item @{ML File.tmp_path}~@{text "path"} relocates the base
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1410
  component of @{text "path"} into the unique temporary directory of
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1411
  the running Isabelle/ML process.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1412
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1413
  \item @{ML serial_string}~@{text "()"} creates a new serial number
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1414
  that is unique over the runtime of the Isabelle/ML process.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1415
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1416
  \end{description}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1417
*}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1418
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1419
text %mlex {* The following example shows how to create unique
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1420
  temporary file names.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1421
*}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1422
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1423
ML {*
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1424
  val tmp1 = File.tmp_path (Path.basic ("foo" ^ serial_string ()));
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1425
  val tmp2 = File.tmp_path (Path.basic ("foo" ^ serial_string ()));
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1426
  @{assert} (tmp1 <> tmp2);
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1427
*}
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1428
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1429
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1430
subsection {* Explicit synchronization *}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1431
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1432
text {* Isabelle/ML also provides some explicit synchronization
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1433
  mechanisms, for the rare situations where mutable shared resources
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1434
  are really required.  These are based on the synchronizations
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1435
  primitives of Poly/ML, which have been adapted to the specific
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1436
  assumptions of the concurrent Isabelle/ML environment.  User code
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1437
  must not use the Poly/ML primitives directly!
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1438
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1439
  \medskip The most basic synchronization concept is a single
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1440
  \emph{critical section} (also called ``monitor'' in the literature).
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1441
  A thread that enters the critical section prevents all other threads
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1442
  from doing the same.  A thread that is already within the critical
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1443
  section may re-enter it in an idempotent manner.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1444
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1445
  Such centralized locking is convenient, because it prevents
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1446
  deadlocks by construction.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1447
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1448
  \medskip More fine-grained locking works via \emph{synchronized
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1449
  variables}.  An explicit state component is associated with
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1450
  mechanisms for locking and signaling.  There are operations to
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1451
  await a condition, change the state, and signal the change to all
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1452
  other waiting threads.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1453
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1454
  Here the synchronized access to the state variable is \emph{not}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1455
  re-entrant: direct or indirect nesting within the same thread will
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1456
  cause a deadlock!
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1457
*}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1458
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1459
text %mlref {*
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1460
  \begin{mldecls}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1461
  @{index_ML NAMED_CRITICAL: "string -> (unit -> 'a) -> 'a"} \\
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1462
  @{index_ML CRITICAL: "(unit -> 'a) -> 'a"} \\
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1463
  \end{mldecls}
39871
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1464
  \begin{mldecls}
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1465
  @{index_ML_type "'a Synchronized.var"} \\
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1466
  @{index_ML Synchronized.var: "string -> 'a -> 'a Synchronized.var"} \\
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1467
  @{index_ML Synchronized.guarded_access: "'a Synchronized.var ->
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1468
  ('a -> ('b * 'a) option) -> 'b"} \\
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1469
  \end{mldecls}
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1470
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1471
  \begin{description}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1472
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1473
  \item @{ML NAMED_CRITICAL}~@{text "name e"} evaluates @{text "e ()"}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1474
  within the central critical section of Isabelle/ML.  No other thread
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1475
  may do so at the same time, but non-critical parallel execution will
39871
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1476
  continue.  The @{text "name"} argument is used for tracing and might
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1477
  help to spot sources of congestion.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1478
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1479
  Entering the critical section without contention is very fast, and
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1480
  several basic system operations do so frequently.  Each thread
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1481
  should leave the critical section quickly, otherwise parallel
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1482
  performance may degrade.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1483
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1484
  \item @{ML CRITICAL} is the same as @{ML NAMED_CRITICAL} with empty
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1485
  name argument.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1486
39871
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1487
  \item Type @{ML_type "'a Synchronized.var"} represents synchronized
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1488
  variables with state of type @{ML_type 'a}.
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1489
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1490
  \item @{ML Synchronized.var}~@{text "name x"} creates a synchronized
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1491
  variable that is initialized with value @{text "x"}.  The @{text
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1492
  "name"} is used for tracing.
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1493
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1494
  \item @{ML Synchronized.guarded_access}~@{text "var f"} lets the
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1495
  function @{text "f"} operate within a critical section on the state
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1496
  @{text "x"} as follows: if @{text "f x"} produces @{ML NONE}, we
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1497
  continue to wait on the internal condition variable, expecting that
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1498
  some other thread will eventually change the content in a suitable
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1499
  manner; if @{text "f x"} produces @{ML SOME}~@{text "(y, x')"} we
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1500
  are satisfied and assign the new state value @{text "x'"}, broadcast
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1501
  a signal to all waiting threads on the associated condition
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1502
  variable, and return the result @{text "y"}.
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1503
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1504
  \end{description}
39871
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1505
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1506
  There are some further variants of the general @{ML
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1507
  Synchronized.guarded_access} combinator, see @{"file"
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1508
  "~~/src/Pure/Concurrent/synchronized.ML"} for details.
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1509
*}
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1510
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1511
text %mlex {* See @{"file" "~~/src/Pure/Concurrent/mailbox.ML"} how to
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1512
  implement a mailbox as synchronized variable over a purely
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1513
  functional queue.
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1514
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1515
  \medskip The following example implements a counter that produces
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1516
  positive integers that are unique over the runtime of the Isabelle
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1517
  process:
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1518
*}
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1519
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1520
ML {*
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1521
  local
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1522
    val counter = Synchronized.var "counter" 0;
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1523
  in
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1524
    fun next () =
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1525
      Synchronized.guarded_access counter
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1526
        (fn i =>
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1527
          let val j = i + 1
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1528
          in SOME (j, j) end);
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1529
  end;
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1530
*}
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1531
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1532
ML {*
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1533
  val a = next ();
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1534
  val b = next ();
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1535
  @{assert} (a <> b);
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1536
*}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1537
30270
61811c9224a6 dummy changes to produce a new changeset of these files;
wenzelm
parents: 29755
diff changeset
  1538
end