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