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