src/Doc/Implementation/ML.thy
author wenzelm
Mon, 23 Jun 2014 12:54:48 +0200
changeset 57350 fc4d65afdf13
parent 57349 4817154180d6
child 57417 29fe9bac501b
permissions -rw-r--r--
more on "Futures"; removed obsolete comments;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
57347
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
     1
(*:wrap=hard:maxLineLen=78:*)
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
     2
29755
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 29646
diff changeset
     3
theory "ML"
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 29646
diff changeset
     4
imports Base
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 29646
diff changeset
     5
begin
18538
88fe84d4d151 outline;
wenzelm
parents: 18537
diff changeset
     6
39822
0de42180febe basic setup for Chapter 0: Isabelle/ML;
wenzelm
parents: 36164
diff changeset
     7
chapter {* Isabelle/ML *}
20489
a684fc70d04e tentative appendix C;
wenzelm
parents: 18554
diff changeset
     8
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
     9
text {* Isabelle/ML is best understood as a certain culture based on
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    10
  Standard ML.  Thus it is not a new programming language, but a
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    11
  certain way to use SML at an advanced level within the Isabelle
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    12
  environment.  This covers a variety of aspects that are geared
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    13
  towards an efficient and robust platform for applications of formal
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    14
  logic with fully foundational proof construction --- according to
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    15
  the well-known \emph{LCF principle}.  There is specific
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    16
  infrastructure with library modules to address the needs of this
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    17
  difficult task.  For example, the raw parallel programming model of
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    18
  Poly/ML is presented as considerably more abstract concept of
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    19
  \emph{future values}, which is then used to augment the inference
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    20
  kernel, proof interpreter, and theory loader accordingly.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    21
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    22
  The main aspects of Isabelle/ML are introduced below.  These
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    23
  first-hand explanations should help to understand how proper
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    24
  Isabelle/ML is to be read and written, and to get access to the
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    25
  wealth of experience that is expressed in the source text and its
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    26
  history of changes.\footnote{See
54703
499f92dc6e45 more antiquotations;
wenzelm
parents: 54387
diff changeset
    27
  @{url "http://isabelle.in.tum.de/repos/isabelle"} for the full
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    28
  Mercurial history.  There are symbolic tags to refer to official
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    29
  Isabelle releases, as opposed to arbitrary \emph{tip} versions that
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    30
  merely reflect snapshots that are never really up-to-date.}  *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    31
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    32
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    33
section {* Style and orthography *}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    34
39879
wenzelm
parents: 39878
diff changeset
    35
text {* The sources of Isabelle/Isar are optimized for
wenzelm
parents: 39878
diff changeset
    36
  \emph{readability} and \emph{maintainability}.  The main purpose is
wenzelm
parents: 39878
diff changeset
    37
  to tell an informed reader what is really going on and how things
wenzelm
parents: 39878
diff changeset
    38
  really work.  This is a non-trivial aim, but it is supported by a
wenzelm
parents: 39878
diff changeset
    39
  certain style of writing Isabelle/ML that has emerged from long
wenzelm
parents: 39878
diff changeset
    40
  years of system development.\footnote{See also the interesting style
wenzelm
parents: 39878
diff changeset
    41
  guide for OCaml
54703
499f92dc6e45 more antiquotations;
wenzelm
parents: 54387
diff changeset
    42
  @{url "http://caml.inria.fr/resources/doc/guides/guidelines.en.html"}
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    43
  which shares many of our means and ends.}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    44
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    45
  The main principle behind any coding style is \emph{consistency}.
39879
wenzelm
parents: 39878
diff changeset
    46
  For a single author of a small program this merely means ``choose
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    47
  your style and stick to it''.  A complex project like Isabelle, with
39879
wenzelm
parents: 39878
diff changeset
    48
  long years of development and different contributors, requires more
wenzelm
parents: 39878
diff changeset
    49
  standardization.  A coding style that is changed every few years or
wenzelm
parents: 39878
diff changeset
    50
  with every new contributor is no style at all, because consistency
wenzelm
parents: 39878
diff changeset
    51
  is quickly lost.  Global consistency is hard to achieve, though.
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    52
  Nonetheless, one should always strive at least for local consistency
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    53
  of modules and sub-systems, without deviating from some general
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    54
  principles how to write Isabelle/ML.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    55
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    56
  In a sense, good coding style is like an \emph{orthography} for the
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    57
  sources: it helps to read quickly over the text and see through the
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    58
  main points, without getting distracted by accidental presentation
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    59
  of free-style code.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    60
*}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    61
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    62
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    63
subsection {* Header and sectioning *}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    64
39879
wenzelm
parents: 39878
diff changeset
    65
text {* Isabelle source files have a certain standardized header
wenzelm
parents: 39878
diff changeset
    66
  format (with precise spacing) that follows ancient traditions
wenzelm
parents: 39878
diff changeset
    67
  reaching back to the earliest versions of the system by Larry
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40508
diff changeset
    68
  Paulson.  See @{file "~~/src/Pure/thm.ML"}, for example.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    69
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    70
  The header includes at least @{verbatim Title} and @{verbatim
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    71
  Author} entries, followed by a prose description of the purpose of
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    72
  the module.  The latter can range from a single line to several
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    73
  paragraphs of explanations.
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    74
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    75
  The rest of the file is divided into sections, subsections,
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    76
  subsubsections, paragraphs etc.\ using a simple layout via ML
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    77
  comments as follows.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    78
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    79
\begin{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    80
(*** section ***)
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    81
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    82
(** subsection **)
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    83
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    84
(* subsubsection *)
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    85
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    86
(*short paragraph*)
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    87
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    88
(*
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    89
  long paragraph,
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    90
  with more text
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    91
*)
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    92
\end{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    93
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    94
  As in regular typography, there is some extra space \emph{before}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    95
  section headings that are adjacent to plain text (not other headings
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    96
  as in the example above).
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    97
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
    98
  \medskip The precise wording of the prose text given in these
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
    99
  headings is chosen carefully to introduce the main theme of the
39879
wenzelm
parents: 39878
diff changeset
   100
  subsequent formal ML text.
wenzelm
parents: 39878
diff changeset
   101
*}
wenzelm
parents: 39878
diff changeset
   102
wenzelm
parents: 39878
diff changeset
   103
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   104
subsection {* Naming conventions *}
39879
wenzelm
parents: 39878
diff changeset
   105
wenzelm
parents: 39878
diff changeset
   106
text {* Since ML is the primary medium to express the meaning of the
wenzelm
parents: 39878
diff changeset
   107
  source text, naming of ML entities requires special care.
wenzelm
parents: 39878
diff changeset
   108
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   109
  \paragraph{Notation.}  A name consists of 1--3 \emph{words} (rarely
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   110
  4, but not more) that are separated by underscore.  There are three
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   111
  variants concerning upper or lower case letters, which are used for
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   112
  certain ML categories as follows:
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   113
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   114
  \medskip
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   115
  \begin{tabular}{lll}
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   116
  variant & example & ML categories \\\hline
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   117
  lower-case & @{ML_text foo_bar} & values, types, record fields \\
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   118
  capitalized & @{ML_text Foo_Bar} & datatype constructors, structures, functors \\
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   119
  upper-case & @{ML_text FOO_BAR} & special values, exception constructors, signatures \\
39880
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,
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   124
  e.g.\ old-style @{ML_text FooBar} instead of @{ML_text Foo_Bar}.
47180
c14fda8fee38 improving spelling
bulwahn
parents: 46262
diff changeset
   125
  Genuine mixed-case names are \emph{not} used, because clear division
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   126
  of words is essential for readability.\footnote{Camel-case was
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   127
  invented to workaround the lack of underscore in some early
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   128
  non-ASCII character sets.  Later it became habitual in some language
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   129
  communities that are now strong in numbers.}
39880
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
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   132
  respect: some Isabelle/ML names are suffixed by extra markers like
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   133
  this: @{ML_text foo_barT}.
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   134
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   135
  Name variants are produced by adding 1--3 primes, e.g.\ @{ML_text
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   136
  foo'}, @{ML_text foo''}, or @{ML_text foo'''}, but not @{ML_text
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   137
  foo''''} or more.  Decimal digits scale better to larger numbers,
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   138
  e.g.\ @{ML_text foo0}, @{ML_text foo1}, @{ML_text 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
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   145
  this is how they are seen in the sources and documentation.  For the
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   146
  same reasons, aliases of well-known library functions should be
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   147
  avoided.
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   148
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   149
  Local names of function abstraction or case/let bindings are
39880
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
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   152
  called @{ML_text helper}, @{ML_text aux}, or @{ML_text f} is
39880
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
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   187
  \paragraph{Specific conventions.} Here are some specific name forms
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   188
  that occur frequently in the sources.
39881
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
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   192
  \item A function that maps @{ML_text foo} to @{ML_text bar} is
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   193
  called @{ML_text foo_to_bar} or @{ML_text bar_of_foo} (never
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   194
  @{ML_text foo2bar}, @{ML_text bar_from_foo}, @{ML_text
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   195
  bar_for_foo}, or @{ML_text bar4foo}).
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   196
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   197
  \item The name component @{ML_text legacy} means that the operation
39881
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
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   200
  \item The name component @{ML_text old} means that this is historic
39881
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
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   203
  \item The name component @{ML_text global} means that this works
39881
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
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   214
  firm naming conventions as follows:
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   215
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   216
  \begin{itemize}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   217
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   218
  \item theories are called @{ML_text thy}, rarely @{ML_text theory}
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   219
  (never @{ML_text thry})
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   220
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   221
  \item proof contexts are called @{ML_text ctxt}, rarely @{ML_text
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   222
  context} (never @{ML_text ctx})
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   223
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   224
  \item generic contexts are called @{ML_text context}, rarely
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   225
  @{ML_text ctxt}
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   226
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   227
  \item local theories are called @{ML_text lthy}, except for local
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   228
  theories that are treated as proof context (which is a semantic
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   229
  super-type)
39881
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
56579
4c94f631c595 tuned spelling;
wenzelm
parents: 56420
diff changeset
   234
  well as semantic prefixes like @{ML_text foo_thy} or @{ML_text
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   235
  bar_ctxt}, but the base conventions above need to be preserved.
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   236
  This allows to visualize the their data flow via plain regular
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   237
  expressions in the editor.
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   238
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   239
  \item The main logical entities (\secref{ch:logic}) have established
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   240
  naming convention as follows:
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   241
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   242
  \begin{itemize}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   243
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   244
  \item sorts are called @{ML_text S}
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   245
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   246
  \item types are called @{ML_text T}, @{ML_text U}, or @{ML_text
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   247
  ty} (never @{ML_text t})
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   248
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   249
  \item terms are called @{ML_text t}, @{ML_text u}, or @{ML_text
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   250
  tm} (never @{ML_text trm})
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   251
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   252
  \item certified types are called @{ML_text cT}, rarely @{ML_text
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   253
  T}, with variants as for types
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   254
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   255
  \item certified terms are called @{ML_text ct}, rarely @{ML_text
52733
wenzelm
parents: 52421
diff changeset
   256
  t}, with variants as for terms (never @{ML_text ctrm})
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   257
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   258
  \item theorems are called @{ML_text th}, or @{ML_text thm}
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   259
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   260
  \end{itemize}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   261
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   262
  Proper semantic names override these conventions completely.  For
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   263
  example, the left-hand side of an equation (as a term) can be called
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   264
  @{ML_text lhs} (not @{ML_text lhs_tm}).  Or a term that is known
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   265
  to be a variable can be called @{ML_text v} or @{ML_text x}.
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   266
40310
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   267
  \item Tactics (\secref{sec:tactics}) are sufficiently important to
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   268
  have specific naming conventions.  The name of a basic tactic
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   269
  definition always has a @{ML_text "_tac"} suffix, the subgoal index
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   270
  (if applicable) is always called @{ML_text i}, and the goal state
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   271
  (if made explicit) is usually called @{ML_text st} instead of the
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   272
  somewhat misleading @{ML_text thm}.  Any other arguments are given
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   273
  before the latter two, and the general context is given first.
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   274
  Example:
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   275
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   276
  \begin{verbatim}
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   277
  fun my_tac ctxt arg1 arg2 i st = ...
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   278
  \end{verbatim}
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   279
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   280
  Note that the goal state @{ML_text st} above is rarely made
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   281
  explicit, if tactic combinators (tacticals) are used as usual.
a0698ec82e6e more on naming tactics;
wenzelm
parents: 40302
diff changeset
   282
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   283
  \end{itemize}
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   284
*}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   285
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   286
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   287
subsection {* General source layout *}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   288
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   289
text {* The general Isabelle/ML source layout imitates regular
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   290
  type-setting to some extent, augmented by the requirements for
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   291
  deeply nested expressions that are commonplace in functional
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   292
  programming.
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   293
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   294
  \paragraph{Line length} is 80 characters according to ancient
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   295
  standards, but we allow as much as 100 characters (not
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   296
  more).\footnote{Readability requires to keep the beginning of a line
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   297
  in view while watching its end.  Modern wide-screen displays do not
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   298
  change the way how the human brain works.  Sources also need to be
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   299
  printable on plain paper with reasonable font-size.} The extra 20
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   300
  characters acknowledge the space requirements due to qualified
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   301
  library references in Isabelle/ML.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   302
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   303
  \paragraph{White-space} is used to emphasize the structure of
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   304
  expressions, following mostly standard conventions for mathematical
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   305
  typesetting, as can be seen in plain {\TeX} or {\LaTeX}.  This
39879
wenzelm
parents: 39878
diff changeset
   306
  defines positioning of spaces for parentheses, punctuation, and
wenzelm
parents: 39878
diff changeset
   307
  infixes as illustrated here:
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   308
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   309
  \begin{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   310
  val x = y + z * (a + b);
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   311
  val pair = (a, b);
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   312
  val record = {foo = 1, bar = 2};
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   313
  \end{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   314
39879
wenzelm
parents: 39878
diff changeset
   315
  Lines are normally broken \emph{after} an infix operator or
wenzelm
parents: 39878
diff changeset
   316
  punctuation character.  For example:
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   317
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   318
  \begin{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   319
  val x =
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   320
    a +
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   321
    b +
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   322
    c;
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   323
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   324
  val tuple =
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   325
   (a,
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   326
    b,
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   327
    c);
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   328
  \end{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   329
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   330
  Some special infixes (e.g.\ @{ML_text "|>"}) work better at the
39879
wenzelm
parents: 39878
diff changeset
   331
  start of the line, but punctuation is always at the end.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   332
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   333
  Function application follows the tradition of @{text "\<lambda>"}-calculus,
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   334
  not informal mathematics.  For example: @{ML_text "f a b"} for a
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   335
  curried function, or @{ML_text "g (a, b)"} for a tupled function.
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   336
  Note that the space between @{ML_text g} and the pair @{ML_text
39879
wenzelm
parents: 39878
diff changeset
   337
  "(a, b)"} follows the important principle of
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   338
  \emph{compositionality}: the layout of @{ML_text "g p"} does not
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   339
  change when @{ML_text "p"} is refined to the concrete pair
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   340
  @{ML_text "(a, b)"}.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   341
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   342
  \paragraph{Indentation} uses plain spaces, never hard
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   343
  tabulators.\footnote{Tabulators were invented to move the carriage
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   344
  of a type-writer to certain predefined positions.  In software they
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   345
  could be used as a primitive run-length compression of consecutive
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   346
  spaces, but the precise result would depend on non-standardized
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   347
  editor configuration.}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   348
39879
wenzelm
parents: 39878
diff changeset
   349
  Each level of nesting is indented by 2 spaces, sometimes 1, very
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   350
  rarely 4, never 8 or any other odd number.
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   351
39879
wenzelm
parents: 39878
diff changeset
   352
  Indentation follows a simple logical format that only depends on the
wenzelm
parents: 39878
diff changeset
   353
  nesting depth, not the accidental length of the text that initiates
wenzelm
parents: 39878
diff changeset
   354
  a level of nesting.  Example:
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   355
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   356
  \begin{verbatim}
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   357
  (* RIGHT *)
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   358
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   359
  if b then
39879
wenzelm
parents: 39878
diff changeset
   360
    expr1_part1
wenzelm
parents: 39878
diff changeset
   361
    expr1_part2
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   362
  else
39879
wenzelm
parents: 39878
diff changeset
   363
    expr2_part1
wenzelm
parents: 39878
diff changeset
   364
    expr2_part2
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   365
39880
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   366
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   367
  (* WRONG *)
7deb6a741626 more on "Naming conventions";
wenzelm
parents: 39879
diff changeset
   368
39879
wenzelm
parents: 39878
diff changeset
   369
  if b then expr1_part1
wenzelm
parents: 39878
diff changeset
   370
            expr1_part2
wenzelm
parents: 39878
diff changeset
   371
  else expr2_part1
wenzelm
parents: 39878
diff changeset
   372
       expr2_part2
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   373
  \end{verbatim}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   374
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   375
  The second form has many problems: it assumes a fixed-width font
39879
wenzelm
parents: 39878
diff changeset
   376
  when viewing the sources, it uses more space on the line and thus
wenzelm
parents: 39878
diff changeset
   377
  makes it hard to observe its strict length limit (working against
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   378
  \emph{readability}), it requires extra editing to adapt the layout
39879
wenzelm
parents: 39878
diff changeset
   379
  to changes of the initial text (working against
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   380
  \emph{maintainability}) etc.
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   381
39879
wenzelm
parents: 39878
diff changeset
   382
  \medskip For similar reasons, any kind of two-dimensional or tabular
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   383
  layouts, ASCII-art with lines or boxes of asterisks etc.\ should be
39879
wenzelm
parents: 39878
diff changeset
   384
  avoided.
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   385
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   386
  \paragraph{Complex expressions} that consist of multi-clausal
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   387
  function definitions, @{ML_text handle}, @{ML_text case},
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   388
  @{ML_text let} (and combinations) require special attention.  The
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   389
  syntax of Standard ML is quite ambitious and admits a lot of
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   390
  variance that can distort the meaning of the text.
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   391
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   392
  Clauses of @{ML_text fun}, @{ML_text fn}, @{ML_text handle},
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   393
  @{ML_text case} get extra indentation to indicate the nesting
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   394
  clearly.  Example:
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   395
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   396
  \begin{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   397
  (* RIGHT *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   398
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   399
  fun foo p1 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   400
        expr1
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   401
    | foo p2 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   402
        expr2
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   403
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   404
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   405
  (* WRONG *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   406
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   407
  fun foo p1 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   408
    expr1
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   409
    | foo p2 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   410
    expr2
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   411
  \end{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   412
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   413
  Body expressions consisting of @{ML_text case} or @{ML_text let}
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   414
  require care to maintain compositionality, to prevent loss of
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   415
  logical indentation where it is especially important to see the
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   416
  structure of the text.  Example:
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   417
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   418
  \begin{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   419
  (* RIGHT *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   420
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   421
  fun foo p1 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   422
        (case e of
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   423
          q1 => ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   424
        | q2 => ...)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   425
    | foo p2 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   426
        let
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   427
          ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   428
        in
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   429
          ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   430
        end
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   431
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   432
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   433
  (* WRONG *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   434
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   435
  fun foo p1 = case e of
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   436
      q1 => ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   437
    | q2 => ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   438
    | foo p2 =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   439
    let
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   440
      ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   441
    in
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   442
      ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   443
    end
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   444
  \end{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   445
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   446
  Extra parentheses around @{ML_text case} expressions are optional,
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   447
  but help to analyse the nesting based on character matching in the
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   448
  editor.
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   449
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   450
  \medskip There are two main exceptions to the overall principle of
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   451
  compositionality in the layout of complex expressions.
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{enumerate}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   454
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   455
  \item @{ML_text "if"} expressions are iterated as if there would be
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   456
  a multi-branch conditional in SML, e.g.
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   457
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   458
  \begin{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   459
  (* RIGHT *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   460
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   461
  if b1 then e1
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   462
  else if b2 then e2
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   463
  else e3
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   464
  \end{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   465
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   466
  \item @{ML_text fn} abstractions are often layed-out as if they
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   467
  would lack any structure by themselves.  This traditional form is
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   468
  motivated by the possibility to shift function arguments back and
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   469
  forth wrt.\ additional combinators.  Example:
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   470
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   471
  \begin{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   472
  (* RIGHT *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   473
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   474
  fun foo x y = fold (fn z =>
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   475
    expr)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   476
  \end{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   477
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   478
  Here the visual appearance is that of three arguments @{ML_text x},
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   479
  @{ML_text y}, @{ML_text z}.
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   480
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   481
  \end{enumerate}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   482
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   483
  Such weakly structured layout should be use with great care.  Here
40153
wenzelm
parents: 40149
diff changeset
   484
  are some counter-examples involving @{ML_text let} expressions:
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   485
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   486
  \begin{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   487
  (* WRONG *)
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   488
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   489
  fun foo x = let
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   490
      val y = ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   491
    in ... end
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   492
41162
2181c47a02fe more correct ML snippets that are unchecked;
wenzelm
parents: 40964
diff changeset
   493
2181c47a02fe more correct ML snippets that are unchecked;
wenzelm
parents: 40964
diff changeset
   494
  (* WRONG *)
2181c47a02fe more correct ML snippets that are unchecked;
wenzelm
parents: 40964
diff changeset
   495
40153
wenzelm
parents: 40149
diff changeset
   496
  fun foo x = let
wenzelm
parents: 40149
diff changeset
   497
    val y = ...
wenzelm
parents: 40149
diff changeset
   498
  in ... end
wenzelm
parents: 40149
diff changeset
   499
41162
2181c47a02fe more correct ML snippets that are unchecked;
wenzelm
parents: 40964
diff changeset
   500
2181c47a02fe more correct ML snippets that are unchecked;
wenzelm
parents: 40964
diff changeset
   501
  (* WRONG *)
2181c47a02fe more correct ML snippets that are unchecked;
wenzelm
parents: 40964
diff changeset
   502
39881
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   503
  fun foo x =
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   504
  let
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   505
    val y = ...
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   506
  in ... end
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   507
  \end{verbatim}
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   508
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   509
  \medskip In general the source layout is meant to emphasize the
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   510
  structure of complex language expressions, not to pretend that SML
40aee0b95ddf more on "Style and orthography";
wenzelm
parents: 39880
diff changeset
   511
  had a completely different syntax (say that of Haskell or Java).
39878
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   512
*}
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   513
31dd361a3060 more on "Style and orthography";
wenzelm
parents: 39875
diff changeset
   514
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   515
section {* SML embedded into Isabelle/Isar *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   516
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   517
text {* ML and Isar are intertwined via an open-ended bootstrap
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   518
  process that provides more and more programming facilities and
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   519
  logical content in an alternating manner.  Bootstrapping starts from
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   520
  the raw environment of existing implementations of Standard ML
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   521
  (mainly Poly/ML, but also SML/NJ).
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   522
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   523
  Isabelle/Pure marks the point where the original ML toplevel is
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   524
  superseded by the Isar toplevel that maintains a uniform context for
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   525
  arbitrary ML values (see also \secref{sec:context}).  This formal
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   526
  environment holds ML compiler bindings, logical entities, and many
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   527
  other things.  Raw SML is never encountered again after the initial
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   528
  bootstrap of Isabelle/Pure.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   529
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   530
  Object-logics like Isabelle/HOL are built within the
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   531
  Isabelle/ML/Isar environment by introducing suitable theories with
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   532
  associated ML modules, either inlined or as separate files.  Thus
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   533
  Isabelle/HOL is defined as a regular user-space application within
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   534
  the Isabelle framework.  Further add-on tools can be implemented in
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   535
  ML within the Isar context in the same manner: ML is part of the
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   536
  standard repertoire of Isabelle, and there is no distinction between
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   537
  ``user'' and ``developer'' in this respect.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   538
*}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   539
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   540
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   541
subsection {* Isar ML commands *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   542
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   543
text {* The primary Isar source language provides facilities to ``open
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   544
  a window'' to the underlying ML compiler.  Especially see the Isar
51295
71fc3776c453 discontinued redundant 'use' command;
wenzelm
parents: 51058
diff changeset
   545
  commands @{command_ref "ML_file"} and @{command_ref "ML"}: both work the
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   546
  same way, only the source text is provided via a file vs.\ inlined,
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   547
  respectively.  Apart from embedding ML into the main theory
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   548
  definition like that, there are many more commands that refer to ML
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   549
  source, such as @{command_ref setup} or @{command_ref declaration}.
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   550
  Even more fine-grained embedding of ML into Isar is encountered in
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   551
  the proof method @{method_ref tactic}, which refines the pending
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   552
  goal state via a given expression of type @{ML_type tactic}.
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   553
*}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   554
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   555
text %mlex {* The following artificial example demonstrates some ML
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   556
  toplevel declarations within the implicit Isar theory context.  This
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   557
  is regular functional programming without referring to logical
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   558
  entities yet.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   559
*}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   560
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   561
ML {*
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   562
  fun factorial 0 = 1
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   563
    | factorial n = n * factorial (n - 1)
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   564
*}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   565
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   566
text {* Here the ML environment is already managed by Isabelle, i.e.\
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   567
  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
   568
  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
   569
  current one in the import hierarchy.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   570
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   571
  Removing the above ML declaration from the source text will remove
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   572
  any trace of this definition as expected.  The Isabelle/ML toplevel
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   573
  environment is managed in a \emph{stateless} way: unlike the raw ML
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   574
  toplevel there are no global side-effects involved
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   575
  here.\footnote{Such a stateless compilation environment is also a
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   576
  prerequisite for robust parallel compilation within independent
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   577
  nodes of the implicit theory development graph.}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   578
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   579
  \medskip The next example shows how to embed ML into Isar proofs, using
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   580
 @{command_ref "ML_prf"} instead of Instead of @{command_ref "ML"}.
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   581
  As illustrated below, the effect on the ML environment is local to
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   582
  the whole proof body, ignoring the block structure.
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   583
*}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   584
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   585
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   586
begin
39851
7219a771ab63 more examples;
wenzelm
parents: 39850
diff changeset
   587
  ML_prf %"ML" {* val a = 1 *}
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   588
  {
39851
7219a771ab63 more examples;
wenzelm
parents: 39850
diff changeset
   589
    ML_prf %"ML" {* val b = a + 1 *}
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   590
  } -- {* Isar block structure ignored by ML environment *}
39851
7219a771ab63 more examples;
wenzelm
parents: 39850
diff changeset
   591
  ML_prf %"ML" {* val c = b + 1 *}
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   592
end
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   593
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   594
text {* By side-stepping the normal scoping rules for Isar proof
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   595
  blocks, embedded ML code can refer to the different contexts and
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   596
  manipulate corresponding entities, e.g.\ export a fact from a block
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   597
  context.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   598
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   599
  \medskip Two further ML commands are useful in certain situations:
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   600
  @{command_ref ML_val} and @{command_ref ML_command} are
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   601
  \emph{diagnostic} in the sense that there is no effect on the
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   602
  underlying environment, and can thus used anywhere (even outside a
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   603
  theory).  The examples below produce long strings of digits by
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   604
  invoking @{ML factorial}: @{command ML_val} already takes care of
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   605
  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
   606
  so we produce an explicit output message.  *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   607
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   608
ML_val {* factorial 100 *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   609
ML_command {* writeln (string_of_int (factorial 100)) *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   610
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   611
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   612
begin
52417
wenzelm
parents: 52416
diff changeset
   613
  ML_val {* factorial 100 *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   614
  ML_command {* writeln (string_of_int (factorial 100)) *}
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   615
end
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   616
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   617
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   618
subsection {* Compile-time context *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   619
39825
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   620
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
   621
  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
   622
  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
   623
  or writing the (local) theory under construction.  Note that such
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   624
  direct access to the compile-time context is rare.  In practice it
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   625
  is typically done via some derived ML functions instead.
39825
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
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   628
text %mlref {*
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   629
  \begin{mldecls}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   630
  @{index_ML ML_Context.the_generic_context: "unit -> Context.generic"} \\
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   631
  @{index_ML "Context.>>": "(Context.generic -> Context.generic) -> unit"} \\
56199
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 55838
diff changeset
   632
  @{index_ML ML_Thms.bind_thms: "string * thm list -> unit"} \\
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 55838
diff changeset
   633
  @{index_ML ML_Thms.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
   634
  \end{mldecls}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   635
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   636
  \begin{description}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   637
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   638
  \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
   639
  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
   640
  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
   641
  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
   642
  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
   643
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   644
  \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
   645
  @{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
   646
56199
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 55838
diff changeset
   647
  \item @{ML ML_Thms.bind_thms}~@{text "(name, thms)"} stores a list of
39850
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   648
  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
   649
  ML toplevel, associating it with the provided name.  Theorems are
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   650
  put into a global ``standard'' format before being stored.
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   651
56199
8e8d28ed7529 tuned signature -- rearranged modules;
wenzelm
parents: 55838
diff changeset
   652
  \item @{ML ML_Thms.bind_thm} is similar to @{ML ML_Thms.bind_thms} but refers to a
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   653
  singleton fact.
39850
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   654
39825
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   655
  \end{description}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   656
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   657
  It is important to note that the above functions are really
39825
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   658
  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
   659
  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
   660
  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
   661
  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
   662
*}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   663
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   664
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   665
subsection {* Antiquotations \label{sec:ML-antiq} *}
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   666
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   667
text {* A very important consequence of embedding SML into Isar is the
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   668
  concept of \emph{ML antiquotation}.  The standard token language of
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   669
  ML is augmented by special syntactic entities of the following form:
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   670
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 54703
diff changeset
   671
  @{rail \<open>
53167
4e7ddd76e632 discontinued unused antiquotation blocks;
wenzelm
parents: 53163
diff changeset
   672
  @{syntax_def antiquote}: '@{' nameref args '}'
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 54703
diff changeset
   673
  \<close>}
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   674
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   675
  Here @{syntax nameref} and @{syntax args} are regular outer syntax
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   676
  categories \cite{isabelle-isar-ref}.  Attributes and proof methods
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   677
  use similar syntax.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   678
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   679
  \medskip A regular antiquotation @{text "@{name args}"} processes
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   680
  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
   681
  produces corresponding ML source text, either as literal
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   682
  \emph{inline} text (e.g. @{text "@{term t}"}) or abstract
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   683
  \emph{value} (e.g. @{text "@{thm th}"}).  This pre-compilation
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   684
  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
   685
  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
   686
  small portions of the code.
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   687
*}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   688
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   689
51636
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   690
subsection {* Printing ML values *}
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   691
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   692
text {* The ML compiler knows about the structure of values according
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   693
  to their static type, and can print them in the manner of the
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   694
  toplevel loop, although the details are non-portable.  The
56399
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   695
  antiquotations @{ML_antiquotation_def "make_string"} and
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   696
  @{ML_antiquotation_def "print"} provide a quasi-portable way to
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   697
  refer to this potential capability of the underlying ML system in
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   698
  generic Isabelle/ML sources.
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   699
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   700
  This is occasionally useful for diagnostic or demonstration
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   701
  purposes.  Note that production-quality tools require proper
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   702
  user-level error messages. *}
51636
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   703
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   704
text %mlantiq {*
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   705
  \begin{matharray}{rcl}
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   706
  @{ML_antiquotation_def "make_string"} & : & @{text ML_antiquotation} \\
56399
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   707
  @{ML_antiquotation_def "print"} & : & @{text ML_antiquotation} \\
51636
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   708
  \end{matharray}
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   709
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 54703
diff changeset
   710
  @{rail \<open>
51636
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   711
  @@{ML_antiquotation make_string}
56399
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   712
  ;
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   713
  @@{ML_antiquotation print} @{syntax name}?
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 54703
diff changeset
   714
  \<close>}
51636
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   715
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   716
  \begin{description}
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   717
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   718
  \item @{text "@{make_string}"} inlines a function to print arbitrary
56399
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   719
  values similar to the ML toplevel.  The result is compiler dependent
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   720
  and may fall back on "?" in certain situations.
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   721
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   722
  \item @{text "@{print f}"} uses the ML function @{text "f: string ->
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   723
  unit"} to output the result of @{text "@{make_string}"} above,
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   724
  together with the source position of the antiquotation.  The default
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   725
  output function is @{ML writeln}.
51636
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   726
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   727
  \end{description}
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   728
*}
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   729
56399
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   730
text %mlex {* The following artificial examples show how to produce
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   731
  adhoc output of ML values for debugging purposes. *}
51636
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   732
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   733
ML {*
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   734
  val x = 42;
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   735
  val y = true;
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   736
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   737
  writeln (@{make_string} {x = x, y = y});
56399
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   738
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   739
  @{print} {x = x, y = y};
386e4cb7ad68 added ML antiquotation @{print};
wenzelm
parents: 56303
diff changeset
   740
  @{print tracing} {x = x, y = y};
51636
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   741
*}
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   742
e49bf0be79ba document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
wenzelm
parents: 51295
diff changeset
   743
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   744
section {* Canonical argument order \label{sec:canonical-argument-order} *}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   745
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   746
text {* Standard ML is a language in the tradition of @{text
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   747
  "\<lambda>"}-calculus and \emph{higher-order functional programming},
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   748
  similar to OCaml, Haskell, or Isabelle/Pure and HOL as logical
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   749
  languages.  Getting acquainted with the native style of representing
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   750
  functions in that setting can save a lot of extra boiler-plate of
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   751
  redundant shuffling of arguments, auxiliary abstractions etc.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   752
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   753
  Functions are usually \emph{curried}: the idea of turning arguments
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   754
  of type @{text "\<tau>\<^sub>i"} (for @{text "i \<in> {1, \<dots> n}"}) into a result of
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   755
  type @{text "\<tau>"} is represented by the iterated function space
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   756
  @{text "\<tau>\<^sub>1 \<rightarrow> \<dots> \<rightarrow> \<tau>\<^sub>n \<rightarrow> \<tau>"}.  This is isomorphic to the well-known
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   757
  encoding via tuples @{text "\<tau>\<^sub>1 \<times> \<dots> \<times> \<tau>\<^sub>n \<rightarrow> \<tau>"}, but the curried
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   758
  version fits more smoothly into the basic calculus.\footnote{The
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   759
  difference is even more significant in higher-order logic, because
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   760
  the redundant tuple structure needs to be accommodated by formal
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   761
  reasoning.}
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   762
56594
e3a06699a13f tuned spelling;
wenzelm
parents: 56579
diff changeset
   763
  Currying gives some flexibility due to \emph{partial application}.  A
53071
wenzelm
parents: 52733
diff changeset
   764
  function @{text "f: \<tau>\<^sub>1 \<rightarrow> \<tau>\<^sub>2 \<rightarrow> \<tau>"} can be applied to @{text "x: \<tau>\<^sub>1"}
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   765
  and the remaining @{text "(f x): \<tau>\<^sub>2 \<rightarrow> \<tau>"} passed to another function
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   766
  etc.  How well this works in practice depends on the order of
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   767
  arguments.  In the worst case, arguments are arranged erratically,
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   768
  and using a function in a certain situation always requires some
56579
4c94f631c595 tuned spelling;
wenzelm
parents: 56420
diff changeset
   769
  glue code.  Thus we would get exponentially many opportunities to
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   770
  decorate the code with meaningless permutations of arguments.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   771
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   772
  This can be avoided by \emph{canonical argument order}, which
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   773
  observes certain standard patterns and minimizes adhoc permutations
40229
wenzelm
parents: 40153
diff changeset
   774
  in their application.  In Isabelle/ML, large portions of text can be
52416
wenzelm
parents: 51636
diff changeset
   775
  written without auxiliary operations like @{text "swap: \<alpha> \<times> \<beta> \<rightarrow> \<beta> \<times>
wenzelm
parents: 51636
diff changeset
   776
  \<alpha>"} or @{text "C: (\<alpha> \<rightarrow> \<beta> \<rightarrow> \<gamma>) \<rightarrow> (\<beta> \<rightarrow> \<alpha> \<rightarrow> \<gamma>)"} (the latter not
wenzelm
parents: 51636
diff changeset
   777
  present in the Isabelle/ML library).
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   778
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   779
  \medskip The basic idea is that arguments that vary less are moved
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   780
  further to the left than those that vary more.  Two particularly
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   781
  important categories of functions are \emph{selectors} and
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   782
  \emph{updates}.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   783
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   784
  The subsequent scheme is based on a hypothetical set-like container
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   785
  of type @{text "\<beta>"} that manages elements of type @{text "\<alpha>"}.  Both
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   786
  the names and types of the associated operations are canonical for
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   787
  Isabelle/ML.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   788
52416
wenzelm
parents: 51636
diff changeset
   789
  \begin{center}
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   790
  \begin{tabular}{ll}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   791
  kind & canonical name and type \\\hline
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   792
  selector & @{text "member: \<beta> \<rightarrow> \<alpha> \<rightarrow> bool"} \\
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   793
  update & @{text "insert: \<alpha> \<rightarrow> \<beta> \<rightarrow> \<beta>"} \\
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   794
  \end{tabular}
52416
wenzelm
parents: 51636
diff changeset
   795
  \end{center}
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   796
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   797
  Given a container @{text "B: \<beta>"}, the partially applied @{text
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   798
  "member B"} is a predicate over elements @{text "\<alpha> \<rightarrow> bool"}, and
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   799
  thus represents the intended denotation directly.  It is customary
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   800
  to pass the abstract predicate to further operations, not the
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   801
  concrete container.  The argument order makes it easy to use other
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   802
  combinators: @{text "forall (member B) list"} will check a list of
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   803
  elements for membership in @{text "B"} etc. Often the explicit
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   804
  @{text "list"} is pointless and can be contracted to @{text "forall
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   805
  (member B)"} to get directly a predicate again.
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   806
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   807
  In contrast, an update operation varies the container, so it moves
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   808
  to the right: @{text "insert a"} is a function @{text "\<beta> \<rightarrow> \<beta>"} to
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   809
  insert a value @{text "a"}.  These can be composed naturally as
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   810
  @{text "insert c \<circ> insert b \<circ> insert a"}.  The slightly awkward
40229
wenzelm
parents: 40153
diff changeset
   811
  inversion of the composition order is due to conventional
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   812
  mathematical notation, which can be easily amended as explained
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   813
  below.
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   814
*}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   815
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   816
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   817
subsection {* Forward application and composition *}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   818
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   819
text {* Regular function application and infix notation works best for
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   820
  relatively deeply structured expressions, e.g.\ @{text "h (f x y + g
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   821
  z)"}.  The important special case of \emph{linear transformation}
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   822
  applies a cascade of functions @{text "f\<^sub>n (\<dots> (f\<^sub>1 x))"}.  This
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   823
  becomes hard to read and maintain if the functions are themselves
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   824
  given as complex expressions.  The notation can be significantly
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   825
  improved by introducing \emph{forward} versions of application and
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   826
  composition as follows:
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   827
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   828
  \medskip
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   829
  \begin{tabular}{lll}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   830
  @{text "x |> f"} & @{text "\<equiv>"} & @{text "f x"} \\
41162
2181c47a02fe more correct ML snippets that are unchecked;
wenzelm
parents: 40964
diff changeset
   831
  @{text "(f #> g) x"} & @{text "\<equiv>"} & @{text "x |> f |> g"} \\
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   832
  \end{tabular}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   833
  \medskip
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   834
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   835
  This enables to write conveniently @{text "x |> f\<^sub>1 |> \<dots> |> f\<^sub>n"} or
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   836
  @{text "f\<^sub>1 #> \<dots> #> f\<^sub>n"} for its functional abstraction over @{text
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   837
  "x"}.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   838
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   839
  \medskip There is an additional set of combinators to accommodate
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   840
  multiple results (via pairs) that are passed on as multiple
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   841
  arguments (via currying).
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   842
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   843
  \medskip
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   844
  \begin{tabular}{lll}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   845
  @{text "(x, y) |-> f"} & @{text "\<equiv>"} & @{text "f x y"} \\
41162
2181c47a02fe more correct ML snippets that are unchecked;
wenzelm
parents: 40964
diff changeset
   846
  @{text "(f #-> g) x"} & @{text "\<equiv>"} & @{text "x |> f |-> g"} \\
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   847
  \end{tabular}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   848
  \medskip
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   849
*}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   850
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   851
text %mlref {*
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   852
  \begin{mldecls}
46262
912b42e64fde tuned ML infixes;
wenzelm
parents: 42665
diff changeset
   853
  @{index_ML_op "|> ": "'a * ('a -> 'b) -> 'b"} \\
912b42e64fde tuned ML infixes;
wenzelm
parents: 42665
diff changeset
   854
  @{index_ML_op "|-> ": "('c * 'a) * ('c -> 'a -> 'b) -> 'b"} \\
912b42e64fde tuned ML infixes;
wenzelm
parents: 42665
diff changeset
   855
  @{index_ML_op "#> ": "('a -> 'b) * ('b -> 'c) -> 'a -> 'c"} \\
912b42e64fde tuned ML infixes;
wenzelm
parents: 42665
diff changeset
   856
  @{index_ML_op "#-> ": "('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd"} \\
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   857
  \end{mldecls}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   858
*}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   859
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   860
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   861
subsection {* Canonical iteration *}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   862
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   863
text {* As explained above, a function @{text "f: \<alpha> \<rightarrow> \<beta> \<rightarrow> \<beta>"} can be
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   864
  understood as update on a configuration of type @{text "\<beta>"},
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   865
  parametrized by arguments of type @{text "\<alpha>"}.  Given @{text "a: \<alpha>"}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   866
  the partial application @{text "(f a): \<beta> \<rightarrow> \<beta>"} operates
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   867
  homogeneously on @{text "\<beta>"}.  This can be iterated naturally over a
53071
wenzelm
parents: 52733
diff changeset
   868
  list of parameters @{text "[a\<^sub>1, \<dots>, a\<^sub>n]"} as @{text "f a\<^sub>1 #> \<dots> #> f a\<^sub>n"}.
wenzelm
parents: 52733
diff changeset
   869
  The latter expression is again a function @{text "\<beta> \<rightarrow> \<beta>"}.
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   870
  It can be applied to an initial configuration @{text "b: \<beta>"} to
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   871
  start the iteration over the given list of arguments: each @{text
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   872
  "a"} in @{text "a\<^sub>1, \<dots>, a\<^sub>n"} is applied consecutively by updating a
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   873
  cumulative configuration.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   874
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   875
  The @{text fold} combinator in Isabelle/ML lifts a function @{text
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   876
  "f"} as above to its iterated version over a list of arguments.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   877
  Lifting can be repeated, e.g.\ @{text "(fold \<circ> fold) f"} iterates
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   878
  over a list of lists as expected.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   879
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   880
  The variant @{text "fold_rev"} works inside-out over the list of
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   881
  arguments, such that @{text "fold_rev f \<equiv> fold f \<circ> rev"} holds.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   882
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   883
  The @{text "fold_map"} combinator essentially performs @{text
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   884
  "fold"} and @{text "map"} simultaneously: each application of @{text
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   885
  "f"} produces an updated configuration together with a side-result;
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   886
  the iteration collects all such side-results as a separate list.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   887
*}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   888
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   889
text %mlref {*
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   890
  \begin{mldecls}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   891
  @{index_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   892
  @{index_ML fold_rev: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   893
  @{index_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   894
  \end{mldecls}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   895
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   896
  \begin{description}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   897
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   898
  \item @{ML fold}~@{text f} lifts the parametrized update function
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   899
  @{text "f"} to a list of parameters.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   900
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   901
  \item @{ML fold_rev}~@{text "f"} is similar to @{ML fold}~@{text
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   902
  "f"}, but works inside-out.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   903
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   904
  \item @{ML fold_map}~@{text "f"} lifts the parametrized update
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   905
  function @{text "f"} (with side-result) to a list of parameters and
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   906
  cumulative side-results.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   907
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   908
  \end{description}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   909
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   910
  \begin{warn}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   911
  The literature on functional programming provides a multitude of
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   912
  combinators called @{text "foldl"}, @{text "foldr"} etc.  SML97
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   913
  provides its own variations as @{ML List.foldl} and @{ML
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   914
  List.foldr}, while the classic Isabelle library also has the
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   915
  historic @{ML Library.foldl} and @{ML Library.foldr}.  To avoid
52416
wenzelm
parents: 51636
diff changeset
   916
  unnecessary complication and confusion, all these historical
wenzelm
parents: 51636
diff changeset
   917
  versions should be ignored, and @{ML fold} (or @{ML fold_rev}) used
wenzelm
parents: 51636
diff changeset
   918
  exclusively.
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   919
  \end{warn}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   920
*}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   921
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   922
text %mlex {* The following example shows how to fill a text buffer
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   923
  incrementally by adding strings, either individually or from a given
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   924
  list.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   925
*}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   926
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   927
ML {*
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   928
  val s =
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   929
    Buffer.empty
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   930
    |> Buffer.add "digits: "
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   931
    |> fold (Buffer.add o string_of_int) (0 upto 9)
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   932
    |> Buffer.content;
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   933
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   934
  @{assert} (s = "digits: 0123456789");
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   935
*}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   936
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   937
text {* Note how @{ML "fold (Buffer.add o string_of_int)"} above saves
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   938
  an extra @{ML "map"} over the given list.  This kind of peephole
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   939
  optimization reduces both the code size and the tree structures in
52416
wenzelm
parents: 51636
diff changeset
   940
  memory (``deforestation''), but it requires some practice to read
wenzelm
parents: 51636
diff changeset
   941
  and write fluently.
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   942
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   943
  \medskip The next example elaborates the idea of canonical
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   944
  iteration, demonstrating fast accumulation of tree content using a
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   945
  text buffer.
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   946
*}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   947
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   948
ML {*
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   949
  datatype tree = Text of string | Elem of string * tree list;
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   950
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   951
  fun slow_content (Text txt) = txt
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   952
    | slow_content (Elem (name, ts)) =
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   953
        "<" ^ name ^ ">" ^
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   954
        implode (map slow_content ts) ^
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   955
        "</" ^ name ^ ">"
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   956
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   957
  fun add_content (Text txt) = Buffer.add txt
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   958
    | add_content (Elem (name, ts)) =
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   959
        Buffer.add ("<" ^ name ^ ">") #>
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   960
        fold add_content ts #>
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   961
        Buffer.add ("</" ^ name ^ ">");
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   962
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   963
  fun fast_content tree =
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   964
    Buffer.empty |> add_content tree |> Buffer.content;
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   965
*}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   966
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   967
text {* The slow part of @{ML slow_content} is the @{ML implode} of
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   968
  the recursive results, because it copies previously produced strings
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   969
  again.
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   970
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   971
  The incremental @{ML add_content} avoids this by operating on a
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
   972
  buffer that is passed through in a linear fashion.  Using @{ML_text
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   973
  "#>"} and contraction over the actual buffer argument saves some
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   974
  additional boiler-plate.  Of course, the two @{ML "Buffer.add"}
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   975
  invocations with concatenated strings could have been split into
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   976
  smaller parts, but this would have obfuscated the source without
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   977
  making a big difference in allocations.  Here we have done some
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   978
  peephole-optimization for the sake of readability.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   979
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   980
  Another benefit of @{ML add_content} is its ``open'' form as a
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   981
  function on buffers that can be continued in further linear
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   982
  transformations, folding etc.  Thus it is more compositional than
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   983
  the naive @{ML slow_content}.  As realistic example, compare the
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   984
  old-style @{ML "Term.maxidx_of_term: term -> int"} with the newer
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   985
  @{ML "Term.maxidx_term: term -> int -> int"} in Isabelle/Pure.
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   986
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   987
  Note that @{ML fast_content} above is only defined as example.  In
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   988
  many practical situations, it is customary to provide the
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   989
  incremental @{ML add_content} only and leave the initialization and
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
   990
  termination to the concrete application by the user.
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   991
*}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   992
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
   993
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   994
section {* Message output channels \label{sec:message-channels} *}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   995
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   996
text {* Isabelle provides output channels for different kinds of
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   997
  messages: regular output, high-volume tracing information, warnings,
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   998
  and errors.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   999
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1000
  Depending on the user interface involved, these messages may appear
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1001
  in different text styles or colours.  The standard output for
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1002
  terminal sessions prefixes each line of warnings by @{verbatim
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1003
  "###"} and errors by @{verbatim "***"}, but leaves anything else
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1004
  unchanged.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1005
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1006
  Messages are associated with the transaction context of the running
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1007
  Isar command.  This enables the front-end to manage commands and
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1008
  resulting messages together.  For example, after deleting a command
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1009
  from a given theory document version, the corresponding message
39872
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1010
  output can be retracted from the display.
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1011
*}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1012
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1013
text %mlref {*
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1014
  \begin{mldecls}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1015
  @{index_ML writeln: "string -> unit"} \\
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1016
  @{index_ML tracing: "string -> unit"} \\
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1017
  @{index_ML warning: "string -> unit"} \\
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1018
  @{index_ML error: "string -> 'a"} \\
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1019
  \end{mldecls}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1020
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1021
  \begin{description}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1022
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1023
  \item @{ML writeln}~@{text "text"} outputs @{text "text"} as regular
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1024
  message.  This is the primary message output operation of Isabelle
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1025
  and should be used by default.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1026
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1027
  \item @{ML tracing}~@{text "text"} outputs @{text "text"} as special
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1028
  tracing message, indicating potential high-volume output to the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1029
  front-end (hundreds or thousands of messages issued by a single
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1030
  command).  The idea is to allow the user-interface to downgrade the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1031
  quality of message display to achieve higher throughput.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1032
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1033
  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
  1034
  output, e.g.\ switch to a different output window.  So this channel
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1035
  should not be used for regular output.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1036
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1037
  \item @{ML warning}~@{text "text"} outputs @{text "text"} as
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1038
  warning, which typically means some extra emphasis on the front-end
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1039
  side (color highlighting, icons, etc.).
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1040
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1041
  \item @{ML error}~@{text "text"} raises exception @{ML ERROR}~@{text
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1042
  "text"} and thus lets the Isar toplevel print @{text "text"} on the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1043
  error channel, which typically means some extra emphasis on the
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1044
  front-end side (color highlighting, icons, etc.).
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1045
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1046
  This assumes that the exception is not handled before the command
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1047
  terminates.  Handling exception @{ML ERROR}~@{text "text"} is a
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1048
  perfectly legal alternative: it means that the error is absorbed
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1049
  without any message output.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1050
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
  1051
  \begin{warn}
54387
890e983cb07b tuned signature;
wenzelm
parents: 53982
diff changeset
  1052
  The actual error channel is accessed via @{ML Output.error_message}, but
51058
98c48d023136 some updates concerning Proof General;
wenzelm
parents: 48985
diff changeset
  1053
  the old interaction protocol of Proof~General \emph{crashes} if that
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1054
  function is used in regular ML code: error output and toplevel
52416
wenzelm
parents: 51636
diff changeset
  1055
  command failure always need to coincide in classic TTY interaction.
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
  1056
  \end{warn}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1057
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
  1058
  \end{description}
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
  1059
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
  1060
  \begin{warn}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1061
  Regular Isabelle/ML code should output messages exclusively by the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1062
  official channels.  Using raw I/O on \emph{stdout} or \emph{stderr}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1063
  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
  1064
  the presence of parallel and asynchronous processing of Isabelle
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1065
  theories.  Such raw output might be displayed by the front-end in
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1066
  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
  1067
  see it.  Moreover, as a genuine side-effect on global process
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1068
  channels, there is no proper way to retract output when Isar command
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1069
  transactions are reset by the system.
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
  1070
  \end{warn}
39872
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1071
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1072
  \begin{warn}
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1073
  The message channels should be used in a message-oriented manner.
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1074
  This means that multi-line output that logically belongs together is
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1075
  issued by a \emph{single} invocation of @{ML writeln} etc.\ with the
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1076
  functional concatenation of all message constituents.
39872
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1077
  \end{warn}
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1078
*}
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1079
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1080
text %mlex {* The following example demonstrates a multi-line
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1081
  warning.  Note that in some situations the user sees only the first
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1082
  line, so the most important point should be made first.
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1083
*}
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1084
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1085
ML_command {*
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1086
  warning (cat_lines
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1087
   ["Beware the Jabberwock, my son!",
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1088
    "The jaws that bite, the claws that catch!",
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1089
    "Beware the Jubjub Bird, and shun",
2b88d00d6790 more on messages;
wenzelm
parents: 39871
diff changeset
  1090
    "The frumious Bandersnatch!"]);
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1091
*}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
  1092
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1093
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1094
section {* Exceptions \label{sec:exceptions} *}
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1095
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1096
text {* The Standard ML semantics of strict functional evaluation
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1097
  together with exceptions is rather well defined, but some delicate
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1098
  points need to be observed to avoid that ML programs go wrong
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1099
  despite static type-checking.  Exceptions in Isabelle/ML are
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1100
  subsequently categorized as follows.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1101
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1102
  \paragraph{Regular user errors.}  These are meant to provide
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1103
  informative feedback about malformed input etc.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1104
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1105
  The \emph{error} function raises the corresponding \emph{ERROR}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1106
  exception, with a plain text message as argument.  \emph{ERROR}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1107
  exceptions can be handled internally, in order to be ignored, turned
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1108
  into other exceptions, or cascaded by appending messages.  If the
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1109
  corresponding Isabelle/Isar command terminates with an \emph{ERROR}
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1110
  exception state, the toplevel will print the result on the error
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1111
  channel (see \secref{sec:message-channels}).
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1112
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1113
  It is considered bad style to refer to internal function names or
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1114
  values in ML source notation in user error messages.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1115
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1116
  Grammatical correctness of error messages can be improved by
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1117
  \emph{omitting} final punctuation: messages are often concatenated
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1118
  or put into a larger context (e.g.\ augmented with source position).
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1119
  By not insisting in the final word at the origin of the error, the
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1120
  system can perform its administrative tasks more easily and
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1121
  robustly.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1122
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1123
  \paragraph{Program failures.}  There is a handful of standard
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1124
  exceptions that indicate general failure situations, or failures of
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1125
  core operations on logical entities (types, terms, theorems,
39856
wenzelm
parents: 39855
diff changeset
  1126
  theories, see \chref{ch:logic}).
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1127
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1128
  These exceptions indicate a genuine breakdown of the program, so the
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1129
  main purpose is to determine quickly what has happened where.
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1130
  Traditionally, the (short) exception message would include the name
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1131
  of an ML function, although this is no longer necessary, because the
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1132
  ML runtime system prints a detailed source position of the
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
  1133
  corresponding @{ML_text raise} keyword.
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1134
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1135
  \medskip User modules can always introduce their own custom
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1136
  exceptions locally, e.g.\ to organize internal failures robustly
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1137
  without overlapping with existing exceptions.  Exceptions that are
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1138
  exposed in module signatures require extra care, though, and should
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1139
  \emph{not} be introduced by default.  Surprise by users of a module
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1140
  can be often minimized by using plain user errors instead.
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1141
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1142
  \paragraph{Interrupts.}  These indicate arbitrary system events:
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1143
  both the ML runtime system and the Isabelle/ML infrastructure signal
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1144
  various exceptional situations by raising the special
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1145
  \emph{Interrupt} exception in user code.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1146
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1147
  This is the one and only way that physical events can intrude an
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1148
  Isabelle/ML program.  Such an interrupt can mean out-of-memory,
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1149
  stack overflow, timeout, internal signaling of threads, or the user
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1150
  producing a console interrupt manually etc.  An Isabelle/ML program
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1151
  that intercepts interrupts becomes dependent on physical effects of
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1152
  the environment.  Even worse, exception handling patterns that are
56579
4c94f631c595 tuned spelling;
wenzelm
parents: 56420
diff changeset
  1153
  too general by accident, e.g.\ by misspelled exception constructors,
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1154
  will cover interrupts unintentionally and thus render the program
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1155
  semantics ill-defined.
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1156
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1157
  Note that the Interrupt exception dates back to the original SML90
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1158
  language definition.  It was excluded from the SML97 version to
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1159
  avoid its malign impact on ML program semantics, but without
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1160
  providing a viable alternative.  Isabelle/ML recovers physical
40229
wenzelm
parents: 40153
diff changeset
  1161
  interruptibility (which is an indispensable tool to implement
wenzelm
parents: 40153
diff changeset
  1162
  managed evaluation of command transactions), but requires user code
wenzelm
parents: 40153
diff changeset
  1163
  to be strictly transparent wrt.\ interrupts.
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1164
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1165
  \begin{warn}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1166
  Isabelle/ML user code needs to terminate promptly on interruption,
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1167
  without guessing at its meaning to the system infrastructure.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1168
  Temporary handling of interrupts for cleanup of global resources
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1169
  etc.\ needs to be followed immediately by re-raising of the original
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1170
  exception.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1171
  \end{warn}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1172
*}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
  1173
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1174
text %mlref {*
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1175
  \begin{mldecls}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1176
  @{index_ML try: "('a -> 'b) -> 'a -> 'b option"} \\
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1177
  @{index_ML can: "('a -> 'b) -> 'a -> bool"} \\
55838
e120a15b0ee6 more antiquotations;
wenzelm
parents: 55837
diff changeset
  1178
  @{index_ML_exception ERROR: string} \\
e120a15b0ee6 more antiquotations;
wenzelm
parents: 55837
diff changeset
  1179
  @{index_ML_exception Fail: string} \\
39856
wenzelm
parents: 39855
diff changeset
  1180
  @{index_ML Exn.is_interrupt: "exn -> bool"} \\
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1181
  @{index_ML reraise: "exn -> 'a"} \\
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56199
diff changeset
  1182
  @{index_ML Runtime.exn_trace: "(unit -> 'a) -> 'a"} \\
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1183
  \end{mldecls}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1184
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1185
  \begin{description}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1186
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1187
  \item @{ML try}~@{text "f x"} makes the partiality of evaluating
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1188
  @{text "f x"} explicit via the option datatype.  Interrupts are
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1189
  \emph{not} handled here, i.e.\ this form serves as safe replacement
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
  1190
  for the \emph{unsafe} version @{ML_text "(SOME"}~@{text "f
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
  1191
  x"}~@{ML_text "handle _ => NONE)"} that is occasionally seen in
52417
wenzelm
parents: 52416
diff changeset
  1192
  books about SML97, not Isabelle/ML.
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1193
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1194
  \item @{ML can} is similar to @{ML try} with more abstract result.
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1195
39856
wenzelm
parents: 39855
diff changeset
  1196
  \item @{ML ERROR}~@{text "msg"} represents user errors; this
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1197
  exception is normally raised indirectly via the @{ML error} function
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1198
  (see \secref{sec:message-channels}).
39856
wenzelm
parents: 39855
diff changeset
  1199
wenzelm
parents: 39855
diff changeset
  1200
  \item @{ML Fail}~@{text "msg"} represents general program failures.
wenzelm
parents: 39855
diff changeset
  1201
wenzelm
parents: 39855
diff changeset
  1202
  \item @{ML Exn.is_interrupt} identifies interrupts robustly, without
wenzelm
parents: 39855
diff changeset
  1203
  mentioning concrete exception constructors in user code.  Handled
wenzelm
parents: 39855
diff changeset
  1204
  interrupts need to be re-raised promptly!
wenzelm
parents: 39855
diff changeset
  1205
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1206
  \item @{ML reraise}~@{text "exn"} raises exception @{text "exn"}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1207
  while preserving its implicit position information (if possible,
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1208
  depending on the ML platform).
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1209
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56199
diff changeset
  1210
  \item @{ML Runtime.exn_trace}~@{ML_text "(fn () =>"}~@{text
40149
4c35be108990 proper markup of uninterpreted ML text as @{ML_text}, not @{verbatim};
wenzelm
parents: 40126
diff changeset
  1211
  "e"}@{ML_text ")"} evaluates expression @{text "e"} while printing
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1212
  a full trace of its stack of nested exceptions (if possible,
53739
599d8c324477 repaired latex (cf. 84522727f9d3);
wenzelm
parents: 53709
diff changeset
  1213
  depending on the ML platform).
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1214
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56199
diff changeset
  1215
  Inserting @{ML Runtime.exn_trace} into ML code temporarily is
53709
84522727f9d3 improved printing of exception trace in Poly/ML 5.5.1;
wenzelm
parents: 53167
diff changeset
  1216
  useful for debugging, but not suitable for production code.
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1217
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1218
  \end{description}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1219
*}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
  1220
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1221
text %mlantiq {*
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1222
  \begin{matharray}{rcl}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1223
  @{ML_antiquotation_def "assert"} & : & @{text ML_antiquotation} \\
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1224
  \end{matharray}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1225
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1226
  \begin{description}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1227
40110
93e7935d4cb5 added ML antiquotation @{assert};
wenzelm
parents: 39883
diff changeset
  1228
  \item @{text "@{assert}"} inlines a function
93e7935d4cb5 added ML antiquotation @{assert};
wenzelm
parents: 39883
diff changeset
  1229
  @{ML_type "bool -> unit"} that raises @{ML Fail} if the argument is
93e7935d4cb5 added ML antiquotation @{assert};
wenzelm
parents: 39883
diff changeset
  1230
  @{ML false}.  Due to inlining the source position of failed
93e7935d4cb5 added ML antiquotation @{assert};
wenzelm
parents: 39883
diff changeset
  1231
  assertions is included in the error output.
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1232
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1233
  \end{description}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1234
*}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1235
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1236
52421
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1237
section {* Strings of symbols \label{sec:symbols} *}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1238
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1239
text {* A \emph{symbol} constitutes the smallest textual unit in
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1240
  Isabelle/ML --- raw ML characters are normally not encountered at
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1241
  all!  Isabelle strings consist of a sequence of symbols, represented
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1242
  as a packed string or an exploded list of strings.  Each symbol is
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1243
  in itself a small string, which has either one of the following
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1244
  forms:
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1245
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1246
  \begin{enumerate}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1247
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1248
  \item a single ASCII character ``@{text "c"}'', for example
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1249
  ``\verb,a,'',
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1250
56579
4c94f631c595 tuned spelling;
wenzelm
parents: 56420
diff changeset
  1251
  \item a codepoint according to UTF-8 (non-ASCII byte sequence),
52421
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1252
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1253
  \item a regular symbol ``\verb,\,\verb,<,@{text "ident"}\verb,>,'',
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1254
  for example ``\verb,\,\verb,<alpha>,'',
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1255
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1256
  \item a control symbol ``\verb,\,\verb,<^,@{text "ident"}\verb,>,'',
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1257
  for example ``\verb,\,\verb,<^bold>,'',
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1258
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1259
  \item a raw symbol ``\verb,\,\verb,<^raw:,@{text text}\verb,>,''
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1260
  where @{text text} consists of printable characters excluding
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1261
  ``\verb,.,'' and ``\verb,>,'', for example
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1262
  ``\verb,\,\verb,<^raw:$\sum_{i = 1}^n$>,'',
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1263
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1264
  \item a numbered raw control symbol ``\verb,\,\verb,<^raw,@{text
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1265
  n}\verb,>, where @{text n} consists of digits, for example
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1266
  ``\verb,\,\verb,<^raw42>,''.
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1267
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1268
  \end{enumerate}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1269
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1270
  The @{text "ident"} syntax for symbol names is @{text "letter
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1271
  (letter | digit)\<^sup>*"}, where @{text "letter = A..Za..z"} and @{text
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1272
  "digit = 0..9"}.  There are infinitely many regular symbols and
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1273
  control symbols, but a fixed collection of standard symbols is
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1274
  treated specifically.  For example, ``\verb,\,\verb,<alpha>,'' is
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1275
  classified as a letter, which means it may occur within regular
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1276
  Isabelle identifiers.
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1277
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1278
  The character set underlying Isabelle symbols is 7-bit ASCII, but
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1279
  8-bit character sequences are passed-through unchanged.  Unicode/UCS
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1280
  data in UTF-8 encoding is processed in a non-strict fashion, such
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1281
  that well-formed code sequences are recognized
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1282
  accordingly.\footnote{Note that ISO-Latin-1 differs from UTF-8 only
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1283
  in some special punctuation characters that even have replacements
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1284
  within the standard collection of Isabelle symbols.  Text consisting
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1285
  of ASCII plus accented letters can be processed in either encoding.}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1286
  Unicode provides its own collection of mathematical symbols, but
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1287
  within the core Isabelle/ML world there is no link to the standard
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1288
  collection of Isabelle regular symbols.
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1289
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1290
  \medskip Output of Isabelle symbols depends on the print mode
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1291
  \cite{isabelle-isar-ref}.  For example, the standard {\LaTeX}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1292
  setup of the Isabelle document preparation system would present
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1293
  ``\verb,\,\verb,<alpha>,'' as @{text "\<alpha>"}, and
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1294
  ``\verb,\,\verb,<^bold>,\verb,\,\verb,<alpha>,'' as @{text "\<^bold>\<alpha>"}.
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1295
  On-screen rendering usually works by mapping a finite subset of
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1296
  Isabelle symbols to suitable Unicode characters.
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1297
*}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1298
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1299
text %mlref {*
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1300
  \begin{mldecls}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1301
  @{index_ML_type "Symbol.symbol": string} \\
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1302
  @{index_ML Symbol.explode: "string -> Symbol.symbol list"} \\
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1303
  @{index_ML Symbol.is_letter: "Symbol.symbol -> bool"} \\
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1304
  @{index_ML Symbol.is_digit: "Symbol.symbol -> bool"} \\
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1305
  @{index_ML Symbol.is_quasi: "Symbol.symbol -> bool"} \\
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1306
  @{index_ML Symbol.is_blank: "Symbol.symbol -> bool"} \\
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1307
  \end{mldecls}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1308
  \begin{mldecls}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1309
  @{index_ML_type "Symbol.sym"} \\
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1310
  @{index_ML Symbol.decode: "Symbol.symbol -> Symbol.sym"} \\
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1311
  \end{mldecls}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1312
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1313
  \begin{description}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1314
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1315
  \item Type @{ML_type "Symbol.symbol"} represents individual Isabelle
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1316
  symbols.
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1317
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1318
  \item @{ML "Symbol.explode"}~@{text "str"} produces a symbol list
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1319
  from the packed form.  This function supersedes @{ML
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1320
  "String.explode"} for virtually all purposes of manipulating text in
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1321
  Isabelle!\footnote{The runtime overhead for exploded strings is
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1322
  mainly that of the list structure: individual symbols that happen to
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1323
  be a singleton string do not require extra memory in Poly/ML.}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1324
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1325
  \item @{ML "Symbol.is_letter"}, @{ML "Symbol.is_digit"}, @{ML
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1326
  "Symbol.is_quasi"}, @{ML "Symbol.is_blank"} classify standard
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1327
  symbols according to fixed syntactic conventions of Isabelle, cf.\
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1328
  \cite{isabelle-isar-ref}.
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1329
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1330
  \item Type @{ML_type "Symbol.sym"} is a concrete datatype that
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1331
  represents the different kinds of symbols explicitly, with
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1332
  constructors @{ML "Symbol.Char"}, @{ML "Symbol.Sym"}, @{ML
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1333
  "Symbol.UTF8"}, @{ML "Symbol.Ctrl"}, @{ML "Symbol.Raw"}.
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1334
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1335
  \item @{ML "Symbol.decode"} converts the string representation of a
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1336
  symbol into the datatype version.
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1337
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1338
  \end{description}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1339
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1340
  \paragraph{Historical note.} In the original SML90 standard the
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1341
  primitive ML type @{ML_type char} did not exists, and @{ML_text
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1342
  "explode: string -> string list"} produced a list of singleton
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1343
  strings like @{ML "raw_explode: string -> string list"} in
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1344
  Isabelle/ML today.  When SML97 came out, Isabelle did not adopt its
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1345
  somewhat anachronistic 8-bit or 16-bit characters, but the idea of
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1346
  exploding a string into a list of small strings was extended to
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1347
  ``symbols'' as explained above.  Thus Isabelle sources can refer to
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1348
  an infinite store of user-defined symbols, without having to worry
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1349
  about the multitude of Unicode encodings that have emerged over the
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1350
  years.  *}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1351
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1352
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1353
section {* Basic data types *}
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1354
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1355
text {* The basis library proposal of SML97 needs to be treated with
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1356
  caution.  Many of its operations simply do not fit with important
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1357
  Isabelle/ML conventions (like ``canonical argument order'', see
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1358
  \secref{sec:canonical-argument-order}), others cause problems with
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1359
  the parallel evaluation model of Isabelle/ML (such as @{ML
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1360
  TextIO.print} or @{ML OS.Process.system}).
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1361
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1362
  Subsequently we give a brief overview of important operations on
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1363
  basic ML data types.
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1364
*}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1365
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1366
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1367
subsection {* Characters *}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1368
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1369
text %mlref {*
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1370
  \begin{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1371
  @{index_ML_type char} \\
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1372
  \end{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1373
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1374
  \begin{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1375
39864
wenzelm
parents: 39863
diff changeset
  1376
  \item Type @{ML_type char} is \emph{not} used.  The smallest textual
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1377
  unit in Isabelle is represented as a ``symbol'' (see
39864
wenzelm
parents: 39863
diff changeset
  1378
  \secref{sec:symbols}).
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1379
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1380
  \end{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1381
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1382
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1383
52421
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1384
subsection {* Strings *}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1385
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1386
text %mlref {*
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1387
  \begin{mldecls}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1388
  @{index_ML_type string} \\
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1389
  \end{mldecls}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1390
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1391
  \begin{description}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1392
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1393
  \item Type @{ML_type string} represents immutable vectors of 8-bit
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1394
  characters.  There are operations in SML to convert back and forth
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1395
  to actual byte vectors, which are seldom used.
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1396
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1397
  This historically important raw text representation is used for
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1398
  Isabelle-specific purposes with the following implicit substructures
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1399
  packed into the string content:
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1400
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1401
  \begin{enumerate}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1402
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1403
  \item sequence of Isabelle symbols (see also \secref{sec:symbols}),
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1404
  with @{ML Symbol.explode} as key operation;
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1405
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1406
  \item XML tree structure via YXML (see also \cite{isabelle-sys}),
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1407
  with @{ML YXML.parse_body} as key operation.
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1408
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1409
  \end{enumerate}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1410
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1411
  Note that Isabelle/ML string literals may refer Isabelle symbols
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1412
  like ``\verb,\,\verb,<alpha>,'' natively, \emph{without} escaping
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1413
  the backslash.  This is a consequence of Isabelle treating all
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1414
  source text as strings of symbols, instead of raw characters.
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1415
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1416
  \end{description}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1417
*}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1418
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1419
text %mlex {* The subsequent example illustrates the difference of
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1420
  physical addressing of bytes versus logical addressing of symbols in
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1421
  Isabelle strings.
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1422
*}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1423
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1424
ML_val {*
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1425
  val s = "\<A>";
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1426
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1427
  @{assert} (length (Symbol.explode s) = 1);
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1428
  @{assert} (size s = 4);
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1429
*}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1430
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1431
text {* Note that in Unicode renderings of the symbol @{text "\<A>"},
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1432
  variations of encodings like UTF-8 or UTF-16 pose delicate questions
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1433
  about the multi-byte representations its codepoint, which is outside
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1434
  of the 16-bit address space of the original Unicode standard from
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1435
  the 1990-ies.  In Isabelle/ML it is just ``\verb,\,\verb,<A>,''
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1436
  literally, using plain ASCII characters beyond any doubts. *}
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1437
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 52420
diff changeset
  1438
39862
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1439
subsection {* Integers *}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1440
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1441
text %mlref {*
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1442
  \begin{mldecls}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1443
  @{index_ML_type int} \\
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1444
  \end{mldecls}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1445
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1446
  \begin{description}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1447
39864
wenzelm
parents: 39863
diff changeset
  1448
  \item Type @{ML_type int} represents regular mathematical integers,
wenzelm
parents: 39863
diff changeset
  1449
  which are \emph{unbounded}.  Overflow never happens in
39862
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1450
  practice.\footnote{The size limit for integer bit patterns in memory
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1451
  is 64\,MB for 32-bit Poly/ML, and much higher for 64-bit systems.}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1452
  This works uniformly for all supported ML platforms (Poly/ML and
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1453
  SML/NJ).
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1454
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1455
  Literal integers in ML text are forced to be of this one true
52417
wenzelm
parents: 52416
diff changeset
  1456
  integer type --- adhoc overloading of SML97 is disabled.
39862
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1457
55837
154855d9a564 clarified names of antiquotations and markup;
wenzelm
parents: 55112
diff changeset
  1458
  Structure @{ML_structure IntInf} of SML97 is obsolete and superseded by
154855d9a564 clarified names of antiquotations and markup;
wenzelm
parents: 55112
diff changeset
  1459
  @{ML_structure Int}.  Structure @{ML_structure Integer} in @{file
39862
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1460
  "~~/src/Pure/General/integer.ML"} provides some additional
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1461
  operations.
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1462
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1463
  \end{description}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1464
*}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1465
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
  1466
40302
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1467
subsection {* Time *}
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1468
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1469
text %mlref {*
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1470
  \begin{mldecls}
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1471
  @{index_ML_type Time.time} \\
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1472
  @{index_ML seconds: "real -> Time.time"} \\
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1473
  \end{mldecls}
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1474
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1475
  \begin{description}
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1476
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1477
  \item Type @{ML_type Time.time} represents time abstractly according
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1478
  to the SML97 basis library definition.  This is adequate for
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1479
  internal ML operations, but awkward in concrete time specifications.
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1480
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1481
  \item @{ML seconds}~@{text "s"} turns the concrete scalar @{text
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1482
  "s"} (measured in seconds) into an abstract time value.  Floating
52417
wenzelm
parents: 52416
diff changeset
  1483
  point numbers are easy to use as configuration options in the
wenzelm
parents: 52416
diff changeset
  1484
  context (see \secref{sec:config-options}) or system preferences that
wenzelm
parents: 52416
diff changeset
  1485
  are maintained externally.
40302
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1486
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1487
  \end{description}
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1488
*}
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1489
2a33038d858b more on "Time" in Isabelle/ML;
wenzelm
parents: 40229
diff changeset
  1490
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1491
subsection {* Options *}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1492
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1493
text %mlref {*
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1494
  \begin{mldecls}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1495
  @{index_ML Option.map: "('a -> 'b) -> 'a option -> 'b option"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1496
  @{index_ML is_some: "'a option -> bool"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1497
  @{index_ML is_none: "'a option -> bool"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1498
  @{index_ML the: "'a option -> 'a"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1499
  @{index_ML these: "'a list option -> 'a list"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1500
  @{index_ML the_list: "'a option -> 'a list"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1501
  @{index_ML the_default: "'a -> 'a option -> 'a"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1502
  \end{mldecls}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1503
*}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1504
52417
wenzelm
parents: 52416
diff changeset
  1505
text {* Apart from @{ML Option.map} most other operations defined in
55837
154855d9a564 clarified names of antiquotations and markup;
wenzelm
parents: 55112
diff changeset
  1506
  structure @{ML_structure Option} are alien to Isabelle/ML an never
52417
wenzelm
parents: 52416
diff changeset
  1507
  used.  The operations shown above are defined in @{file
wenzelm
parents: 52416
diff changeset
  1508
  "~~/src/Pure/General/basics.ML"}.  *}
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1509
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1510
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1511
subsection {* Lists *}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1512
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1513
text {* Lists are ubiquitous in ML as simple and light-weight
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1514
  ``collections'' for many everyday programming tasks.  Isabelle/ML
39874
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1515
  provides important additions and improvements over operations that
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1516
  are predefined in the SML97 library. *}
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1517
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1518
text %mlref {*
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1519
  \begin{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1520
  @{index_ML cons: "'a -> 'a list -> 'a list"} \\
39874
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1521
  @{index_ML member: "('b * 'a -> bool) -> 'a list -> 'b -> bool"} \\
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1522
  @{index_ML insert: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1523
  @{index_ML remove: "('b * 'a -> bool) -> 'b -> 'a list -> 'a list"} \\
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1524
  @{index_ML update: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1525
  \end{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1526
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1527
  \begin{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1528
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1529
  \item @{ML cons}~@{text "x xs"} evaluates to @{text "x :: xs"}.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1530
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1531
  Tupled infix operators are a historical accident in Standard ML.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1532
  The curried @{ML cons} amends this, but it should be only used when
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1533
  partial application is required.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1534
39874
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1535
  \item @{ML member}, @{ML insert}, @{ML remove}, @{ML update} treat
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1536
  lists as a set-like container that maintains the order of elements.
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40508
diff changeset
  1537
  See @{file "~~/src/Pure/library.ML"} for the full specifications
39874
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1538
  (written in ML).  There are some further derived operations like
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1539
  @{ML union} or @{ML inter}.
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1540
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1541
  Note that @{ML insert} is conservative about elements that are
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1542
  already a @{ML member} of the list, while @{ML update} ensures that
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1543
  the latest entry is always put in front.  The latter discipline is
39874
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1544
  often more appropriate in declarations of context data
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1545
  (\secref{sec:context-data}) that are issued by the user in Isar
52417
wenzelm
parents: 52416
diff changeset
  1546
  source: later declarations take precedence over earlier ones.
39874
bbac63bbcffe clarified "lists as a set-like container";
wenzelm
parents: 39872
diff changeset
  1547
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1548
  \end{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1549
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1550
52417
wenzelm
parents: 52416
diff changeset
  1551
text %mlex {* Using canonical @{ML fold} together with @{ML cons} (or
wenzelm
parents: 52416
diff changeset
  1552
  similar standard operations) alternates the orientation of data.
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1553
  The is quite natural and should not be altered forcible by inserting
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1554
  extra applications of @{ML rev}.  The alternative @{ML fold_rev} can
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1555
  be used in the few situations, where alternation should be
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1556
  prevented.
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1557
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1558
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1559
ML {*
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1560
  val items = [1, 2, 3, 4, 5, 6, 7, 8, 9, 10];
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1561
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1562
  val list1 = fold cons items [];
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1563
  @{assert} (list1 = rev items);
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1564
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1565
  val list2 = fold_rev cons items [];
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
  1566
  @{assert} (list2 = items);
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1567
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1568
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1569
text {* The subsequent example demonstrates how to \emph{merge} two
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1570
  lists in a natural way. *}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1571
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1572
ML {*
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1573
  fun merge_lists eq (xs, ys) = fold_rev (insert eq) ys xs;
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1574
*}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1575
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1576
text {* Here the first list is treated conservatively: only the new
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1577
  elements from the second list are inserted.  The inside-out order of
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1578
  insertion via @{ML fold_rev} attempts to preserve the order of
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1579
  elements in the result.
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1580
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1581
  This way of merging lists is typical for context data
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1582
  (\secref{sec:context-data}).  See also @{ML merge} as defined in
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40508
diff changeset
  1583
  @{file "~~/src/Pure/library.ML"}.
39883
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1584
*}
3d3d6038bdaa more on "Canonical argument order";
wenzelm
parents: 39881
diff changeset
  1585
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
  1586
39875
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1587
subsection {* Association lists *}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1588
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1589
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
  1590
  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
  1591
  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
  1592
  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
  1593
  occurrence into account.
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1594
*}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1595
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1596
text {*
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1597
  \begin{mldecls}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1598
  @{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
  1599
  @{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
  1600
  @{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
  1601
  \end{mldecls}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1602
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1603
  \begin{description}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1604
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1605
  \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
  1606
  implement the main ``framework operations'' for mappings in
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1607
  Isabelle/ML, following standard conventions for their names and
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1608
  types.
39875
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1609
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1610
  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
  1611
  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
  1612
  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
  1613
  @{text "the_element"} or @{text "get"}.
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1614
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1615
  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
  1616
  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
  1617
  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
  1618
  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
  1619
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1620
  \end{description}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1621
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1622
  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
  1623
  implementation of finite mappings in many practical situations.  A
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40508
diff changeset
  1624
  more heavy-duty table structure is defined in @{file
39875
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1625
  "~~/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
  1626
  thousands or millions of elements.
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1627
*}
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1628
648c930125f6 more on "Association lists", based on more succinct version of older material;
wenzelm
parents: 39874
diff changeset
  1629
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1630
subsection {* Unsynchronized references *}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1631
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1632
text %mlref {*
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1633
  \begin{mldecls}
39870
wenzelm
parents: 39868
diff changeset
  1634
  @{index_ML_type "'a Unsynchronized.ref"} \\
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1635
  @{index_ML Unsynchronized.ref: "'a -> 'a Unsynchronized.ref"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1636
  @{index_ML "!": "'a Unsynchronized.ref -> 'a"} \\
46262
912b42e64fde tuned ML infixes;
wenzelm
parents: 42665
diff changeset
  1637
  @{index_ML_op ":=": "'a Unsynchronized.ref * 'a -> unit"} \\
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1638
  \end{mldecls}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1639
*}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1640
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1641
text {* Due to ubiquitous parallelism in Isabelle/ML (see also
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1642
  \secref{sec:multi-threading}), the mutable reference cells of
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1643
  Standard ML are notorious for causing problems.  In a highly
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1644
  parallel system, both correctness \emph{and} performance are easily
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1645
  degraded when using mutable data.
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1646
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1647
  The unwieldy name of @{ML Unsynchronized.ref} for the constructor
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1648
  for references in Isabelle/ML emphasizes the inconveniences caused by
46262
912b42e64fde tuned ML infixes;
wenzelm
parents: 42665
diff changeset
  1649
  mutability.  Existing operations @{ML "!"}  and @{ML_op ":="} are
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1650
  unchanged, but should be used with special precautions, say in a
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1651
  strictly local situation that is guaranteed to be restricted to
40508
76894f975440 never open Unsynchronized;
wenzelm
parents: 40310
diff changeset
  1652
  sequential evaluation --- now and in the future.
76894f975440 never open Unsynchronized;
wenzelm
parents: 40310
diff changeset
  1653
76894f975440 never open Unsynchronized;
wenzelm
parents: 40310
diff changeset
  1654
  \begin{warn}
76894f975440 never open Unsynchronized;
wenzelm
parents: 40310
diff changeset
  1655
  Never @{ML_text "open Unsynchronized"}, not even in a local scope!
76894f975440 never open Unsynchronized;
wenzelm
parents: 40310
diff changeset
  1656
  Pretending that mutable state is no problem is a very bad idea.
76894f975440 never open Unsynchronized;
wenzelm
parents: 40310
diff changeset
  1657
  \end{warn}
76894f975440 never open Unsynchronized;
wenzelm
parents: 40310
diff changeset
  1658
*}
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
  1659
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1660
39870
wenzelm
parents: 39868
diff changeset
  1661
section {* Thread-safe programming \label{sec:multi-threading} *}
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1662
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1663
text {* Multi-threaded execution has become an everyday reality in
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1664
  Isabelle since Poly/ML 5.2.1 and Isabelle2008.  Isabelle/ML provides
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1665
  implicit and explicit parallelism by default, and there is no way
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1666
  for user-space tools to ``opt out''.  ML programs that are purely
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1667
  functional, output messages only via the official channels
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1668
  (\secref{sec:message-channels}), and do not intercept interrupts
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1669
  (\secref{sec:exceptions}) can participate in the multi-threaded
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1670
  environment immediately without further ado.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1671
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1672
  More ambitious tools with more fine-grained interaction with the
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1673
  environment need to observe the principles explained below.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1674
*}
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1675
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1676
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1677
subsection {* Multi-threading with shared memory *}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1678
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1679
text {* Multiple threads help to organize advanced operations of the
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1680
  system, such as real-time conditions on command transactions,
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1681
  sub-components with explicit communication, general asynchronous
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1682
  interaction etc.  Moreover, parallel evaluation is a prerequisite to
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1683
  make adequate use of the CPU resources that are available on
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1684
  multi-core systems.\footnote{Multi-core computing does not mean that
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1685
  there are ``spare cycles'' to be wasted.  It means that the
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1686
  continued exponential speedup of CPU performance due to ``Moore's
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1687
  Law'' follows different rules: clock frequency has reached its peak
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1688
  around 2005, and applications need to be parallelized in order to
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1689
  avoid a perceived loss of performance.  See also
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1690
  \cite{Sutter:2005}.}
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1691
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1692
  Isabelle/Isar exploits the inherent structure of theories and proofs
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1693
  to support \emph{implicit parallelism} to a large extent.  LCF-style
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1694
  theorem provides almost ideal conditions for that, see also
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1695
  \cite{Wenzel:2009}.  This means, significant parts of theory and
52418
f00e4d8dde4c updated to Isabelle2013;
wenzelm
parents: 52417
diff changeset
  1696
  proof checking is parallelized by default.  In Isabelle2013, a
f00e4d8dde4c updated to Isabelle2013;
wenzelm
parents: 52417
diff changeset
  1697
  maximum speedup-factor of 3.5 on 4 cores and 6.5 on 8 cores can be
f00e4d8dde4c updated to Isabelle2013;
wenzelm
parents: 52417
diff changeset
  1698
  expected.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1699
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1700
  \medskip ML threads lack the memory protection of separate
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1701
  processes, and operate concurrently on shared heap memory.  This has
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1702
  the advantage that results of independent computations are directly
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1703
  available to other threads: abstract values can be passed without
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1704
  copying or awkward serialization that is typically required for
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1705
  separate processes.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1706
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1707
  To make shared-memory multi-threading work robustly and efficiently,
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1708
  some programming guidelines need to be observed.  While the ML
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1709
  system is responsible to maintain basic integrity of the
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1710
  representation of ML values in memory, the application programmer
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1711
  needs to ensure that multi-threaded execution does not break the
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1712
  intended semantics.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1713
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1714
  \begin{warn}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1715
  To participate in implicit parallelism, tools need to be
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1716
  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
  1717
  performance of the whole system.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1718
  \end{warn}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1719
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1720
  Apart from observing the principles of thread-safeness passively,
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1721
  advanced tools may also exploit parallelism actively, e.g.\ by using
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1722
  ``future values'' (\secref{sec:futures}) or the more basic library
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1723
  functions for parallel list operations (\secref{sec:parlist}).
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1724
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1725
  \begin{warn}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1726
  Parallel computing resources are managed centrally by the
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1727
  Isabelle/ML infrastructure.  User programs must not fork their own
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1728
  ML threads to perform computations.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1729
  \end{warn}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1730
*}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1731
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1732
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1733
subsection {* Critical shared resources *}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1734
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1735
text {* Thread-safeness is mainly concerned about concurrent
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1736
  read/write access to shared resources, which are outside the purely
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1737
  functional world of ML.  This covers the following in particular.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1738
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1739
  \begin{itemize}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1740
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1741
  \item Global references (or arrays), i.e.\ mutable memory cells that
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1742
  persist over several invocations of associated
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1743
  operations.\footnote{This is independent of the visibility of such
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1744
  mutable values in the toplevel scope.}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1745
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1746
  \item Global state of the running Isabelle/ML process, i.e.\ raw I/O
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1747
  channels, environment variables, current working directory.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1748
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1749
  \item Writable resources in the file-system that are shared among
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1750
  different threads or external processes.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1751
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1752
  \end{itemize}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1753
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1754
  Isabelle/ML provides various mechanisms to avoid critical shared
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1755
  resources in most situations.  As last resort there are some
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1756
  mechanisms for explicit synchronization.  The following guidelines
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1757
  help to make Isabelle/ML programs work smoothly in a concurrent
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1758
  environment.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1759
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1760
  \begin{itemize}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1761
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1762
  \item Avoid global references altogether.  Isabelle/Isar maintains a
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1763
  uniform context that incorporates arbitrary data declared by user
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1764
  programs (\secref{sec:context-data}).  This context is passed as
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1765
  plain value and user tools can get/map their own data in a purely
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1766
  functional manner.  Configuration options within the context
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1767
  (\secref{sec:config-options}) provide simple drop-in replacements
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1768
  for historic reference variables.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1769
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1770
  \item Keep components with local state information re-entrant.
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1771
  Instead of poking initial values into (private) global references, a
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1772
  new state record can be created on each invocation, and passed
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1773
  through any auxiliary functions of the component.  The state record
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1774
  may well contain mutable references, without requiring any special
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1775
  synchronizations, as long as each invocation gets its own copy.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1776
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1777
  \item Avoid raw output on @{text "stdout"} or @{text "stderr"}.  The
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1778
  Poly/ML library is thread-safe for each individual output operation,
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1779
  but the ordering of parallel invocations is arbitrary.  This means
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1780
  raw output will appear on some system console with unpredictable
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1781
  interleaving of atomic chunks.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1782
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1783
  Note that this does not affect regular message output channels
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1784
  (\secref{sec:message-channels}).  An official message is associated
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1785
  with the command transaction from where it originates, independently
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1786
  of other transactions.  This means each running Isar command has
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1787
  effectively its own set of message channels, and interleaving can
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1788
  only happen when commands use parallelism internally (and only at
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1789
  message boundaries).
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1790
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1791
  \item Treat environment variables and the current working directory
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1792
  of the running process as strictly read-only.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1793
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1794
  \item Restrict writing to the file-system to unique temporary files.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1795
  Isabelle already provides a temporary directory that is unique for
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1796
  the running process, and there is a centralized source of unique
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1797
  serial numbers in Isabelle/ML.  Thus temporary files that are passed
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1798
  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
  1799
  thread-safe.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1800
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1801
  \end{itemize}
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1802
*}
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1803
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1804
text %mlref {*
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1805
  \begin{mldecls}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1806
  @{index_ML File.tmp_path: "Path.T -> Path.T"} \\
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1807
  @{index_ML serial_string: "unit -> string"} \\
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1808
  \end{mldecls}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1809
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1810
  \begin{description}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1811
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1812
  \item @{ML File.tmp_path}~@{text "path"} relocates the base
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1813
  component of @{text "path"} into the unique temporary directory of
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1814
  the running Isabelle/ML process.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1815
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1816
  \item @{ML serial_string}~@{text "()"} creates a new serial number
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1817
  that is unique over the runtime of the Isabelle/ML process.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1818
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1819
  \end{description}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1820
*}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1821
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1822
text %mlex {* The following example shows how to create unique
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1823
  temporary file names.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1824
*}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1825
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1826
ML {*
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1827
  val tmp1 = File.tmp_path (Path.basic ("foo" ^ serial_string ()));
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1828
  val tmp2 = File.tmp_path (Path.basic ("foo" ^ serial_string ()));
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1829
  @{assert} (tmp1 <> tmp2);
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1830
*}
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1831
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1832
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1833
subsection {* Explicit synchronization *}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1834
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1835
text {* Isabelle/ML also provides some explicit synchronization
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1836
  mechanisms, for the rare situations where mutable shared resources
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1837
  are really required.  These are based on the synchronizations
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1838
  primitives of Poly/ML, which have been adapted to the specific
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1839
  assumptions of the concurrent Isabelle/ML environment.  User code
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1840
  must not use the Poly/ML primitives directly!
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1841
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1842
  \medskip The most basic synchronization concept is a single
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1843
  \emph{critical section} (also called ``monitor'' in the literature).
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1844
  A thread that enters the critical section prevents all other threads
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1845
  from doing the same.  A thread that is already within the critical
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1846
  section may re-enter it in an idempotent manner.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1847
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1848
  Such centralized locking is convenient, because it prevents
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1849
  deadlocks by construction.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1850
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1851
  \medskip More fine-grained locking works via \emph{synchronized
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1852
  variables}.  An explicit state component is associated with
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1853
  mechanisms for locking and signaling.  There are operations to
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1854
  await a condition, change the state, and signal the change to all
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1855
  other waiting threads.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1856
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1857
  Here the synchronized access to the state variable is \emph{not}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1858
  re-entrant: direct or indirect nesting within the same thread will
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1859
  cause a deadlock!
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1860
*}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1861
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1862
text %mlref {*
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1863
  \begin{mldecls}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1864
  @{index_ML NAMED_CRITICAL: "string -> (unit -> 'a) -> 'a"} \\
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1865
  @{index_ML CRITICAL: "(unit -> 'a) -> 'a"} \\
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1866
  \end{mldecls}
39871
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1867
  \begin{mldecls}
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1868
  @{index_ML_type "'a Synchronized.var"} \\
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1869
  @{index_ML Synchronized.var: "string -> 'a -> 'a Synchronized.var"} \\
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1870
  @{index_ML Synchronized.guarded_access: "'a Synchronized.var ->
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1871
  ('a -> ('b * 'a) option) -> 'b"} \\
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1872
  \end{mldecls}
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1873
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1874
  \begin{description}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1875
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1876
  \item @{ML NAMED_CRITICAL}~@{text "name e"} evaluates @{text "e ()"}
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1877
  within the central critical section of Isabelle/ML.  No other thread
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1878
  may do so at the same time, but non-critical parallel execution will
39871
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1879
  continue.  The @{text "name"} argument is used for tracing and might
39868
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1880
  help to spot sources of congestion.
732ab20fec3b misc tuning;
wenzelm
parents: 39867
diff changeset
  1881
52418
f00e4d8dde4c updated to Isabelle2013;
wenzelm
parents: 52417
diff changeset
  1882
  Entering the critical section without contention is very fast.  Each
f00e4d8dde4c updated to Isabelle2013;
wenzelm
parents: 52417
diff changeset
  1883
  thread should stay within the critical section only very briefly,
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1884
  otherwise parallel performance may degrade.
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1885
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1886
  \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
  1887
  name argument.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1888
39871
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1889
  \item Type @{ML_type "'a Synchronized.var"} represents synchronized
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1890
  variables with state of type @{ML_type 'a}.
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1891
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1892
  \item @{ML Synchronized.var}~@{text "name x"} creates a synchronized
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1893
  variable that is initialized with value @{text "x"}.  The @{text
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1894
  "name"} is used for tracing.
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1895
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1896
  \item @{ML Synchronized.guarded_access}~@{text "var f"} lets the
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1897
  function @{text "f"} operate within a critical section on the state
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1898
  @{text "x"} as follows: if @{text "f x"} produces @{ML NONE}, it
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1899
  continues to wait on the internal condition variable, expecting that
39871
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1900
  some other thread will eventually change the content in a suitable
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1901
  manner; if @{text "f x"} produces @{ML SOME}~@{text "(y, x')"} it is
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1902
  satisfied and assigns the new state value @{text "x'"}, broadcasts a
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1903
  signal to all waiting threads on the associated condition variable,
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1904
  and returns the result @{text "y"}.
39871
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1905
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1906
  \end{description}
39871
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1907
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1908
  There are some further variants of the @{ML
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40508
diff changeset
  1909
  Synchronized.guarded_access} combinator, see @{file
39871
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1910
  "~~/src/Pure/Concurrent/synchronized.ML"} for details.
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1911
*}
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1912
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1913
text %mlex {* The following example implements a counter that produces
39871
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1914
  positive integers that are unique over the runtime of the Isabelle
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1915
  process:
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1916
*}
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1917
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1918
ML {*
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1919
  local
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1920
    val counter = Synchronized.var "counter" 0;
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1921
  in
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1922
    fun next () =
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1923
      Synchronized.guarded_access counter
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1924
        (fn i =>
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1925
          let val j = i + 1
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1926
          in SOME (j, j) end);
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1927
  end;
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1928
*}
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1929
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1930
ML {*
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1931
  val a = next ();
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1932
  val b = next ();
b905971407a1 more on synchronized variables;
wenzelm
parents: 39870
diff changeset
  1933
  @{assert} (a <> b);
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1934
*}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
  1935
40800
330eb65c9469 Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents: 40508
diff changeset
  1936
text {* \medskip See @{file "~~/src/Pure/Concurrent/mailbox.ML"} how
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1937
  to implement a mailbox as synchronized variable over a purely
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1938
  functional queue. *}
916cb4a28ffd misc tuning;
wenzelm
parents: 40110
diff changeset
  1939
52419
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1940
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1941
section {* Managed evaluation *}
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1942
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1943
text {* Execution of Standard ML follows the model of strict
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1944
  functional evaluation with optional exceptions.  Evaluation happens
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1945
  whenever some function is applied to (sufficiently many)
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1946
  arguments. The result is either an explicit value or an implicit
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1947
  exception.
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1948
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1949
  \emph{Managed evaluation} in Isabelle/ML organizes expressions and
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1950
  results to control certain physical side-conditions, to say more
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1951
  specifically when and how evaluation happens.  For example, the
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1952
  Isabelle/ML library supports lazy evaluation with memoing, parallel
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1953
  evaluation via futures, asynchronous evaluation via promises,
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1954
  evaluation with time limit etc.
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1955
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1956
  \medskip An \emph{unevaluated expression} is represented either as
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1957
  unit abstraction @{verbatim "fn () => a"} of type
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1958
  @{verbatim "unit -> 'a"} or as regular function
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1959
  @{verbatim "fn a => b"} of type @{verbatim "'a -> 'b"}.  Both forms
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1960
  occur routinely, and special care is required to tell them apart ---
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1961
  the static type-system of SML is only of limited help here.
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1962
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1963
  The first form is more intuitive: some combinator @{text "(unit ->
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1964
  'a) -> 'a"} applies the given function to @{text "()"} to initiate
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1965
  the postponed evaluation process.  The second form is more flexible:
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1966
  some combinator @{text "('a -> 'b) -> 'a -> 'b"} acts like a
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1967
  modified form of function application; several such combinators may
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1968
  be cascaded to modify a given function, before it is ultimately
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1969
  applied to some argument.
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1970
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1971
  \medskip \emph{Reified results} make the disjoint sum of regular
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1972
  values versions exceptional situations explicit as ML datatype:
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1973
  @{text "'a result = Res of 'a | Exn of exn"}.  This is typically
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1974
  used for administrative purposes, to store the overall outcome of an
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1975
  evaluation process.
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1976
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1977
  \emph{Parallel exceptions} aggregate reified results, such that
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1978
  multiple exceptions are digested as a collection in canonical form
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1979
  that identifies exceptions according to their original occurrence.
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1980
  This is particular important for parallel evaluation via futures
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1981
  \secref{sec:futures}, which are organized as acyclic graph of
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1982
  evaluations that depend on other evaluations: exceptions stemming
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1983
  from shared sub-graphs are exposed exactly once and in the order of
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1984
  their original occurrence (e.g.\ when printed at the toplevel).
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1985
  Interrupt counts as neutral element here: it is treated as minimal
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1986
  information about some canceled evaluation process, and is absorbed
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1987
  by the presence of regular program exceptions.  *}
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1988
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1989
text %mlref {*
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1990
  \begin{mldecls}
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1991
  @{index_ML_type "'a Exn.result"} \\
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1992
  @{index_ML Exn.capture: "('a -> 'b) -> 'a -> 'b Exn.result"} \\
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1993
  @{index_ML Exn.interruptible_capture: "('a -> 'b) -> 'a -> 'b Exn.result"} \\
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1994
  @{index_ML Exn.release: "'a Exn.result -> 'a"} \\
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1995
  @{index_ML Par_Exn.release_all: "'a Exn.result list -> 'a list"} \\
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1996
  @{index_ML Par_Exn.release_first: "'a Exn.result list -> 'a list"} \\
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1997
  \end{mldecls}
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1998
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  1999
  \begin{description}
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2000
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2001
  \item Type @{ML_type "'a Exn.result"} represents the disjoint sum of
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2002
  ML results explicitly, with constructor @{ML Exn.Res} for regular
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2003
  values and @{ML "Exn.Exn"} for exceptions.
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2004
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2005
  \item @{ML Exn.capture}~@{text "f x"} manages the evaluation of
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2006
  @{text "f x"} such that exceptions are made explicit as @{ML
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2007
  "Exn.Exn"}.  Note that this includes physical interrupts (see also
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2008
  \secref{sec:exceptions}), so the same precautions apply to user
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2009
  code: interrupts must not be absorbed accidentally!
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2010
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2011
  \item @{ML Exn.interruptible_capture} is similar to @{ML
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2012
  Exn.capture}, but interrupts are immediately re-raised as required
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2013
  for user code.
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2014
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2015
  \item @{ML Exn.release}~@{text "result"} releases the original
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2016
  runtime result, exposing its regular value or raising the reified
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2017
  exception.
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2018
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2019
  \item @{ML Par_Exn.release_all}~@{text "results"} combines results
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2020
  that were produced independently (e.g.\ by parallel evaluation).  If
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2021
  all results are regular values, that list is returned.  Otherwise,
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2022
  the collection of all exceptions is raised, wrapped-up as collective
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2023
  parallel exception.  Note that the latter prevents access to
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2024
  individual exceptions by conventional @{verbatim "handle"} of SML.
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2025
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2026
  \item @{ML Par_Exn.release_first} is similar to @{ML
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2027
  Par_Exn.release_all}, but only the first exception that has occurred
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2028
  in the original evaluation process is raised again, the others are
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2029
  ignored.  That single exception may get handled by conventional
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2030
  means in SML.
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2031
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2032
  \end{description}
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2033
*}
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2034
24018d1167a3 more on managed evaluation;
wenzelm
parents: 52418
diff changeset
  2035
52420
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2036
subsection {* Parallel skeletons \label{sec:parlist} *}
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2037
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2038
text {*
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2039
  Algorithmic skeletons are combinators that operate on lists in
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2040
  parallel, in the manner of well-known @{text map}, @{text exists},
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2041
  @{text forall} etc.  Management of futures (\secref{sec:futures})
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2042
  and their results as reified exceptions is wrapped up into simple
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2043
  programming interfaces that resemble the sequential versions.
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2044
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2045
  What remains is the application-specific problem to present
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2046
  expressions with suitable \emph{granularity}: each list element
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2047
  corresponds to one evaluation task.  If the granularity is too
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2048
  coarse, the available CPUs are not saturated.  If it is too
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2049
  fine-grained, CPU cycles are wasted due to the overhead of
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2050
  organizing parallel processing.  In the worst case, parallel
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2051
  performance will be less than the sequential counterpart!
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2052
*}
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2053
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2054
text %mlref {*
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2055
  \begin{mldecls}
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2056
  @{index_ML Par_List.map: "('a -> 'b) -> 'a list -> 'b list"} \\
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2057
  @{index_ML Par_List.get_some: "('a -> 'b option) -> 'a list -> 'b option"} \\
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2058
  \end{mldecls}
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2059
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2060
  \begin{description}
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2061
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2062
  \item @{ML Par_List.map}~@{text "f [x\<^sub>1, \<dots>, x\<^sub>n]"} is like @{ML
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2063
  "map"}~@{text "f [x\<^sub>1, \<dots>, x\<^sub>n]"}, but the evaluation of @{text "f x\<^sub>i"}
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2064
  for @{text "i = 1, \<dots>, n"} is performed in parallel.
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2065
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2066
  An exception in any @{text "f x\<^sub>i"} cancels the overall evaluation
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2067
  process.  The final result is produced via @{ML
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2068
  Par_Exn.release_first} as explained above, which means the first
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2069
  program exception that happened to occur in the parallel evaluation
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2070
  is propagated, and all other failures are ignored.
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2071
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2072
  \item @{ML Par_List.get_some}~@{text "f [x\<^sub>1, \<dots>, x\<^sub>n]"} produces some
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2073
  @{text "f x\<^sub>i"} that is of the form @{text "SOME y\<^sub>i"}, if that
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2074
  exists, otherwise @{text "NONE"}.  Thus it is similar to @{ML
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2075
  Library.get_first}, but subject to a non-deterministic parallel
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2076
  choice process.  The first successful result cancels the overall
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2077
  evaluation process; other exceptions are propagated as for @{ML
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2078
  Par_List.map}.
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2079
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2080
  This generic parallel choice combinator is the basis for derived
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2081
  forms, such as @{ML Par_List.find_some}, @{ML Par_List.exists}, @{ML
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2082
  Par_List.forall}.
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2083
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2084
  \end{description}
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2085
*}
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2086
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2087
text %mlex {* Subsequently, the Ackermann function is evaluated in
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2088
  parallel for some ranges of arguments. *}
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2089
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2090
ML_val {*
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2091
  fun ackermann 0 n = n + 1
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2092
    | ackermann m 0 = ackermann (m - 1) 1
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2093
    | ackermann m n = ackermann (m - 1) (ackermann m (n - 1));
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2094
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2095
  Par_List.map (ackermann 2) (500 upto 1000);
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2096
  Par_List.map (ackermann 3) (5 upto 10);
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2097
*}
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2098
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2099
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2100
subsection {* Lazy evaluation *}
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2101
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2102
text {*
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2103
  Classic lazy evaluation works via the @{text lazy}~/ @{text force} pair of
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2104
  operations: @{text lazy} to wrap an unevaluated expression, and @{text
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2105
  force} to evaluate it once and store its result persistently. Later
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2106
  invocations of @{text force} retrieve the stored result without another
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2107
  evaluation. Isabelle/ML refines this idea to accommodate the aspects of
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2108
  multi-threading, synchronous program exceptions and asynchronous interrupts.
57347
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2109
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2110
  The first thread that invokes @{text force} on an unfinished lazy value
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2111
  changes its state into a \emph{promise} of the eventual result and starts
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2112
  evaluating it. Any other threads that @{text force} the same lazy value in
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2113
  the meantime need to wait for it to finish, by producing a regular result or
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2114
  program exception. If the evaluation attempt is interrupted, this event is
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2115
  propagated to all waiting threads and the lazy value is reset to its
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2116
  original state.
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2117
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2118
  This means a lazy value is completely evaluated at most once, in a
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2119
  thread-safe manner. There might be multiple interrupted evaluation attempts,
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2120
  and multiple receivers of intermediate interrupt events. Interrupts are
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2121
  \emph{not} made persistent: later evaluation attempts start again from the
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2122
  original expression.
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2123
*}
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2124
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2125
text %mlref {*
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2126
  \begin{mldecls}
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2127
  @{index_ML_type "'a lazy"} \\
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2128
  @{index_ML Lazy.lazy: "(unit -> 'a) -> 'a lazy"} \\
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2129
  @{index_ML Lazy.value: "'a -> 'a lazy"} \\
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2130
  @{index_ML Lazy.force: "'a lazy -> 'a"} \\
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2131
  \end{mldecls}
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2132
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2133
  \begin{description}
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2134
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2135
  \item Type @{ML_type "'a lazy"} represents lazy values over type @{verbatim
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2136
  "'a"}.
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2137
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2138
  \item @{ML Lazy.lazy}~@{text "(fn () => e)"} wraps the unevaluated
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2139
  expression @{text e} as unfinished lazy value.
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2140
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2141
  \item @{ML Lazy.value}~@{text a} wraps the value @{text a} as finished lazy
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2142
  value.  When forced, it returns @{text a} without any further evaluation.
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2143
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2144
  There is very low overhead for this proforma wrapping of strict values as
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2145
  lazy values.
57347
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2146
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2147
  \item @{ML Lazy.force}~@{text x} produces the result of the lazy value in a
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2148
  thread-safe manner as explained above. Thus it may cause the current thread
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2149
  to wait on a pending evaluation attempt by another thread.
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2150
5c8f4a35d8d7 more on "Lazy evaluation";
wenzelm
parents: 56594
diff changeset
  2151
  \end{description}
52420
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2152
*}
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2153
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2154
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2155
subsection {* Futures \label{sec:futures} *}
52420
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2156
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2157
text {*
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2158
  Futures help to organize parallel execution in a value-oriented manner, with
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2159
  @{text fork}~/ @{text join} as the main pair of operations, and some further
57350
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2160
  variants; see also \cite{Wenzel:2009,Wenzel:2013:ITP}. Unlike lazy values,
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2161
  futures are evaluated strictly and spontaneously on separate worker threads.
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2162
  Futures may be canceled, which leads to interrupts on running evaluation
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2163
  attempts, and forces structurally related futures to fail for all time;
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2164
  already finished futures remain unchanged. Exceptions between related
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2165
  futures are propagated as well, and turned into parallel exceptions (see
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2166
  above).
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2167
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2168
  Technically, a future is a single-assignment variable together with a
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2169
  \emph{task} that serves administrative purposes, notably within the
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2170
  \emph{task queue} where new futures are registered for eventual evaluation
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2171
  and the worker threads retrieve their work.
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2172
57350
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2173
  The pool of worker threads is limited, in correlation with the number of
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2174
  physical cores on the machine. Note that allocation of runtime resources may
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2175
  be distorted either if workers yield CPU time (e.g.\ via system sleep or
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2176
  wait operations), or if non-worker threads contend for significant runtime
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2177
  resources independently. There is a limited number of replacement worker
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2178
  threads that get activated in certain explicit wait conditions, after a
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2179
  timeout.
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2180
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2181
  \medskip Each future task belongs to some \emph{task group}, which
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2182
  represents the hierarchic structure of related tasks, together with the
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2183
  exception status a that point. By default, the task group of a newly created
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2184
  future is a new sub-group of the presently running one, but it is also
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2185
  possible to indicate different group layouts under program control.
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2186
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2187
  Cancellation of futures actually refers to the corresponding task group and
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2188
  all its sub-groups. Thus interrupts are propagated down the group hierarchy.
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2189
  Regular program exceptions are treated likewise: failure of the evaluation
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2190
  of some future task affects its own group and all sub-groups. Given a
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2191
  particular task group, its \emph{group status} cumulates all relevant
57350
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2192
  exceptions according to its position within the group hierarchy. Interrupted
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2193
  tasks that lack regular result information, will pick up parallel exceptions
fc4d65afdf13 more on "Futures";
wenzelm
parents: 57349
diff changeset
  2194
  from the cumulative group status.
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2195
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2196
  \medskip A \emph{passive future} or \emph{promise} is a future with slightly
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2197
  different evaluation policies: there is only a single-assignment variable
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2198
  and some expression to evaluate for the \emph{failed} case (e.g.\ to clean
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2199
  up resources when canceled). A regular result is produced by external means,
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2200
  using a separate \emph{fulfill} operation.
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2201
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2202
  Promises are managed in the same task queue, so regular futures may depend
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2203
  on them. This allows a form of reactive programming, where some promises are
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2204
  used as minimal elements (or guards) within the future dependency graph:
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2205
  when these promises are fulfilled the evaluation of subsequent futures
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2206
  starts spontaneously, according to their own inter-dependencies.
52420
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2207
*}
81d5072935ac more on managed evaluation;
wenzelm
parents: 52419
diff changeset
  2208
57348
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2209
text %mlref {*
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2210
  \begin{mldecls}
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2211
  @{index_ML_type "'a future"} \\
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2212
  @{index_ML Future.fork: "(unit -> 'a) -> 'a future"} \\
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2213
  @{index_ML Future.forks: "Future.params -> (unit -> 'a) list -> 'a future list"} \\
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2214
  @{index_ML Future.join: "'a future -> 'a"} \\
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2215
  @{index_ML Future.joins: "'a future list -> 'a list"} \\
57348
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2216
  @{index_ML Future.value: "'a -> 'a future"} \\
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2217
  @{index_ML Future.map: "('a -> 'b) -> 'a future -> 'b future"} \\
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2218
  @{index_ML Future.cancel: "'a future -> unit"} \\
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2219
  @{index_ML Future.cancel_group: "Future.group -> unit"} \\[0.5ex]
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2220
  @{index_ML Future.promise: "(unit -> unit) -> 'a future"} \\
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2221
  @{index_ML Future.fulfill: "'a future -> 'a -> unit"} \\
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2222
  \end{mldecls}
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2223
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2224
  \begin{description}
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2225
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2226
  \item Type @{ML_type "'a future"} represents future values over type
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2227
  @{verbatim "'a"}.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2228
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2229
  \item @{ML Future.fork}~@{text "(fn () => e)"} registers the unevaluated
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2230
  expression @{text e} as unfinished future value, to be evaluated eventually
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2231
  on the parallel worker-thread farm. This is a shorthand for @{ML
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2232
  Future.forks} below, with default parameters and a single expression.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2233
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2234
  \item @{ML Future.forks}~@{text "params exprs"} is the general interface to
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2235
  fork several futures simultaneously. The @{text params} consist of the
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2236
  following fields:
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2237
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2238
  \begin{itemize}
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2239
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2240
  \item @{text "name : string"} (default @{ML "\"\""}) specifies a common name
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2241
  for the tasks of the forked futures, which serves diagnostic purposes.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2242
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2243
  \item @{text "group : Future.group option"} (default @{ML NONE}) specifies
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2244
  an optional task group for the forked futures. @{ML NONE} means that a new
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2245
  sub-group of the current worker-thread task context is created. If this is
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2246
  not a worker thread, the group will be a new root in the group hierarchy.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2247
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2248
  \item @{text "deps : Future.task list"} (default @{ML "[]"}) specifies
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2249
  dependencies on other future tasks, i.e.\ the adjacency relation in the
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2250
  global task queue. Dependencies on already finished tasks are ignored.
57348
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2251
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2252
  \item @{text "pri : int"} (default @{ML 0}) specifies a priority within the
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2253
  task queue.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2254
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2255
  Typically there is only little deviation from the default priority @{ML 0}.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2256
  As a rule of thumb, @{ML "~1"} means ``low priority" and @{ML 1} means
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2257
  ``high priority''.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2258
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2259
  Note that the task priority only affects the position in the queue, not the
57348
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2260
  thread priority. When a worker thread picks up a task for processing, it
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2261
  runs with the normal thread priority to the end (or until canceled). Higher
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2262
  priority tasks that are queued later need to wait until this (or another)
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2263
  worker thread becomes free again.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2264
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2265
  \item @{text "interrupts : bool"} (default @{ML true}) tells whether the
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2266
  worker thread that processes the corresponding task is initially put into
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2267
  interruptible state. This state may change again while running, by modifying
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2268
  the thread attributes.
57348
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2269
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2270
  With interrupts disabled, a running future task cannot be canceled.  It is
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2271
  the responsibility of the programmer that this special state is retained
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2272
  only briefly.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2273
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2274
  \end{itemize}
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2275
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2276
  \item @{ML Future.join}~@{text x} retrieves the value of an already finished
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2277
  future, which may lead to an exception, according to the result of its
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2278
  previous evaluation.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2279
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2280
  For an unfinished future there are several cases depending on the role of
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2281
  the current thread and the status of the future. A non-worker thread waits
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2282
  passively until the future is eventually evaluated. A worker thread
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2283
  temporarily changes its task context and takes over the responsibility to
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2284
  evaluate the future expression on the spot. The latter is done in a
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2285
  thread-safe manner: other threads that intend to join the same future need
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2286
  to wait until the ongoing evaluation is finished.
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2287
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2288
  Note that excessive use of dynamic dependencies of futures by adhoc joining
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2289
  may lead to bad utilization of CPU cores, due to threads waiting on other
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2290
  threads to finish required futures. The future task farm has a limited
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2291
  amount of replacement threads that continue working on unrelated tasks after
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2292
  some timeout.
57348
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2293
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2294
  Whenever possible, static dependencies of futures should be specified
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2295
  explicitly when forked (see @{text deps} above). Thus the evaluation can
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2296
  work from the bottom up, without join conflicts and wait states.
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2297
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2298
  \item @{ML Future.joins}~@{text xs} joins the given list of futures
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2299
  simultaneously, which is more efficient than @{ML "map Future.join"}~@{text
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2300
  xs}.
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2301
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2302
  Based on the dependency graph of tasks, the current thread takes over the
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2303
  responsibility to evaluate future expressions that are required for the main
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2304
  result, working from the bottom up. Waiting on future results that are
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2305
  presently evaluated on other threads only happens as last resort, when no
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2306
  other unfinished futures are left over.
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2307
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2308
  \item @{ML Future.value}~@{text a} wraps the value @{text a} as finished
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2309
  future value, bypassing the worker-thread farm. When joined, it returns
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2310
  @{text a} without any further evaluation.
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2311
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2312
  There is very low overhead for this proforma wrapping of strict values as
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2313
  future values.
57348
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2314
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2315
  \item @{ML Future.map}~@{text "f x"} is a fast-path implementation of @{ML
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2316
  Future.fork}~@{text "(fn () => f ("}@{ML Future.join}~@{text "x))"}, which
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2317
  avoids the full overhead of the task queue and worker-thread farm as far as
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2318
  possible. The function @{text f} is supposed to be some trivial
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2319
  post-processing or projection of the future result.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2320
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2321
  \item @{ML Future.cancel}~@{text "x"} cancels the task group of the given
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2322
  future, using @{ML Future.cancel_group} below.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2323
57349
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2324
  \item @{ML Future.cancel_group}~@{text "group"} cancels all tasks of the
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2325
  given task group for all time. Threads that are presently processing a task
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2326
  of the given group are interrupted: it may take some time until they are
4817154180d6 more on "Futures";
wenzelm
parents: 57348
diff changeset
  2327
  actually terminated. Tasks that are queued but not yet processed are
57348
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2328
  dequeued and forced into interrupted state. Since the task group is itself
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2329
  invalidated, any further attempt to fork a future that belongs to it will
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2330
  yield a canceled result as well.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2331
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2332
  \item @{ML Future.promise}~@{text abort} registers a passive future with the
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2333
  given @{text abort} operation: it is invoked when the future task group is
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2334
  canceled.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2335
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2336
  \item @{ML Future.fulfill}~@{text "x a"} finishes the passive future @{text
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2337
  x} by the given value @{text a}. If the promise has already been canceled,
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2338
  the attempt to fulfill it causes an exception.
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2339
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2340
  \end{description}
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2341
*}
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2342
a6b1847b6335 more on "Future values";
wenzelm
parents: 57347
diff changeset
  2343
47180
c14fda8fee38 improving spelling
bulwahn
parents: 46262
diff changeset
  2344
end