doc-src/IsarImplementation/Thy/ML.thy
author wenzelm
Mon, 18 Oct 2010 21:37:26 +0100
changeset 39867 a8363532cd4d
parent 39866 5ec01d5acd0c
child 39868 732ab20fec3b
permissions -rw-r--r--
somewhat modernized version of "Thread-safe programming";
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29755
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 29646
diff changeset
     1
theory "ML"
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 29646
diff changeset
     2
imports Base
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 29646
diff changeset
     3
begin
18538
88fe84d4d151 outline;
wenzelm
parents: 18537
diff changeset
     4
39822
0de42180febe basic setup for Chapter 0: Isabelle/ML;
wenzelm
parents: 36164
diff changeset
     5
chapter {* Isabelle/ML *}
20489
a684fc70d04e tentative appendix C;
wenzelm
parents: 18554
diff changeset
     6
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
     7
text {* Isabelle/ML is best understood as a certain culture based on
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
     8
  Standard ML.  Thus it is not a new programming language, but a
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
     9
  certain way to use SML at an advanced level within the Isabelle
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    10
  environment.  This covers a variety of aspects that are geared
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    11
  towards an efficient and robust platform for applications of formal
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    12
  logic with fully foundational proof construction --- according to
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    13
  the well-known \emph{LCF principle}.  There are specific library
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    14
  modules and infrastructure to address the needs for such difficult
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    15
  tasks.  For example, the raw parallel programming model of Poly/ML
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    16
  is presented as considerably more abstract concept of \emph{future
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    17
  values}, which is then used to augment the inference kernel, proof
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    18
  interpreter, and theory loader accordingly.
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    19
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    20
  The main aspects of Isabelle/ML are introduced below.  These
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    21
  first-hand explanations should help to understand how proper
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    22
  Isabelle/ML is to be read and written, and to get access to the
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    23
  wealth of experience that is expressed in the source text and its
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    24
  history of changes.\footnote{See
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    25
  \url{http://isabelle.in.tum.de/repos/isabelle} for the full
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    26
  Mercurial history.  There are symbolic tags to refer to official
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    27
  Isabelle releases, as opposed to arbitrary \emph{tip} versions that
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    28
  merely reflect snapshots that are never really up-to-date.}  *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    29
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    30
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    31
section {* SML embedded into Isabelle/Isar *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    32
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    33
text {* ML and Isar are intertwined via an open-ended bootstrap
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    34
  process that provides more and more programming facilities and
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    35
  logical content in an alternating manner.  Bootstrapping starts from
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    36
  the raw environment of existing implementations of Standard ML
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    37
  (mainly Poly/ML, but also SML/NJ).
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    38
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    39
  Isabelle/Pure marks the point where the original ML toplevel is
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    40
  superseded by the Isar toplevel that maintains a uniform environment
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    41
  for arbitrary ML values (see also \secref{sec:context}).  This
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    42
  formal context holds logical entities as well as ML compiler
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    43
  bindings, among many other things.  Raw Standard ML is never
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    44
  encountered again after the initial bootstrap of Isabelle/Pure.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    45
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    46
  Object-logics such as Isabelle/HOL are built within the
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    47
  Isabelle/ML/Isar environment of Pure by introducing suitable
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    48
  theories with associated ML text, either inlined or as separate
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    49
  files.  Thus Isabelle/HOL is defined as a regular user-space
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    50
  application within the Isabelle framework.  Further add-on tools can
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    51
  be implemented in ML within the Isar context in the same manner: ML
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    52
  is part of the regular user-space repertoire of Isabelle.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    53
*}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    54
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    55
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
    56
subsection {* Isar ML commands *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    57
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    58
text {* The primary Isar source language provides various facilities
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    59
  to open a ``window'' to the underlying ML compiler.  Especially see
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    60
  @{command_ref "use"} and @{command_ref "ML"}, which work exactly the
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    61
  same way, only the source text is provided via a file vs.\ inlined,
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    62
  respectively.  Apart from embedding ML into the main theory
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    63
  definition like that, there are many more commands that refer to ML
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    64
  source, such as @{command_ref setup} or @{command_ref declaration}.
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    65
  An example of even more fine-grained embedding of ML into Isar is
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    66
  the proof method @{method_ref tactic}, which refines the pending goal state
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    67
  via a given expression of type @{ML_type tactic}.
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    68
*}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    69
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    70
text %mlex {* The following artificial example demonstrates some ML
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    71
  toplevel declarations within the implicit Isar theory context.  This
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    72
  is regular functional programming without referring to logical
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    73
  entities yet.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    74
*}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    75
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    76
ML {*
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    77
  fun factorial 0 = 1
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    78
    | factorial n = n * factorial (n - 1)
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    79
*}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    80
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
    81
text {* Here the ML environment is really managed by Isabelle, i.e.\
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
    82
  the @{ML factorial} function is not yet accessible in the preceding
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
    83
  paragraph, nor in a different theory that is independent from the
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
    84
  current one in the import hierarchy.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    85
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    86
  Removing the above ML declaration from the source text will remove
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    87
  any trace of this definition as expected.  The Isabelle/ML toplevel
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    88
  environment is managed in a \emph{stateless} way: unlike the raw ML
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    89
  toplevel or similar command loops of Computer Algebra systems, there
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    90
  are no global side-effects involved here.\footnote{Such a stateless
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    91
  compilation environment is also a prerequisite for robust parallel
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    92
  compilation within independent nodes of the implicit theory
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    93
  development graph.}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
    94
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
    95
  \medskip The next example shows how to embed ML into Isar proofs.
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    96
  Instead of @{command_ref "ML"} for theory mode, we use @{command_ref
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    97
  "ML_prf"} for proof mode.  As illustrated below, its effect on the
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
    98
  ML environment is local to the whole proof body, while ignoring its
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
    99
  internal block structure.  *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   100
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   101
example_proof
39851
7219a771ab63 more examples;
wenzelm
parents: 39850
diff changeset
   102
  ML_prf %"ML" {* val a = 1 *}
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   103
  { -- {* Isar block structure ignored by ML environment *}
39851
7219a771ab63 more examples;
wenzelm
parents: 39850
diff changeset
   104
    ML_prf %"ML" {* val b = a + 1 *}
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   105
  } -- {* Isar block structure ignored by ML environment *}
39851
7219a771ab63 more examples;
wenzelm
parents: 39850
diff changeset
   106
  ML_prf %"ML" {* val c = b + 1 *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   107
qed
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   108
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   109
text {* By side-stepping the normal scoping rules for Isar proof
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   110
  blocks, embedded ML code can refer to the different contexts
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   111
  explicitly, and manipulate corresponding entities, e.g.\ export a
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   112
  fact from a block context.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   113
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   114
  \medskip Two further ML commands are useful in certain situations:
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   115
  @{command_ref ML_val} and @{command_ref ML_command} are both
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   116
  \emph{diagnostic} in the sense that there is no effect on the
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   117
  underlying environment, and can thus used anywhere (even outside a
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   118
  theory).  The examples below produce long strings of digits by
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   119
  invoking @{ML factorial}: @{command ML_val} already takes care of
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   120
  printing the ML toplevel result, but @{command ML_command} is silent
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   121
  so we produce an explicit output message.  *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   122
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   123
ML_val {* factorial 100 *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   124
ML_command {* writeln (string_of_int (factorial 100)) *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   125
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   126
example_proof
39824
679075565542 misc tuning;
wenzelm
parents: 39823
diff changeset
   127
  ML_val {* factorial 100 *}  (* FIXME check/fix indentation *)
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   128
  ML_command {* writeln (string_of_int (factorial 100)) *}
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   129
qed
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   130
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   131
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   132
subsection {* Compile-time context *}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   133
39825
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   134
text {* Whenever the ML compiler is invoked within Isabelle/Isar, the
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   135
  formal context is passed as a thread-local reference variable.  Thus
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   136
  ML code may access the theory context during compilation, by reading
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   137
  or writing the (local) theory under construction.  Note that such
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   138
  direct access to the compile-time context is rare; in practice it is
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   139
  typically via some derived ML functions.
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   140
*}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   141
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   142
text %mlref {*
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   143
  \begin{mldecls}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   144
  @{index_ML ML_Context.the_generic_context: "unit -> Context.generic"} \\
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   145
  @{index_ML "Context.>> ": "(Context.generic -> Context.generic) -> unit"} \\
39850
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   146
  @{index_ML bind_thms: "string * thm list -> unit"} \\
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   147
  @{index_ML bind_thm: "string * thm -> unit"} \\
39825
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   148
  \end{mldecls}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   149
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   150
  \begin{description}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   151
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   152
  \item @{ML "ML_Context.the_generic_context ()"} refers to the theory
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   153
  context of the ML toplevel --- at compile time.  ML code needs to
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   154
  take care to refer to @{ML "ML_Context.the_generic_context ()"}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   155
  correctly.  Recall that evaluation of a function body is delayed
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   156
  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
   157
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   158
  \item @{ML "Context.>>"}~@{text f} applies context transformation
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   159
  @{text f} to the implicit context of the ML toplevel.
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   160
39850
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   161
  \item @{ML bind_thms}~@{text "(name, thms)"} stores a list of
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   162
  theorems produced in ML both in the (global) theory context and the
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   163
  ML toplevel, associating it with the provided name.  Theorems are
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   164
  put into a global ``standard'' format before being stored.
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   165
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   166
  \item @{ML bind_thm} is similar to @{ML bind_thms} but refers to a
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   167
  singleton theorem.
f4c614ece7ed moved bind_thm(s) to implementation manual;
wenzelm
parents: 39835
diff changeset
   168
39825
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   169
  \end{description}
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   170
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   171
  It is very important to note that the above functions are really
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   172
  restricted to the compile time, even though the ML compiler is
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   173
  invoked at run-time.  The majority of ML code either uses static
39825
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   174
  antiquotations (\secref{sec:ML-antiq}) or refers to the theory or
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   175
  proof context at run-time, by explicit functional abstraction.
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39824
diff changeset
   176
*}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   177
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   178
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   179
subsection {* Antiquotations \label{sec:ML-antiq} *}
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   180
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   181
text {* A very important consequence of embedding SML into Isar is the
39829
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   182
  concept of \emph{ML antiquotation}.  First, the standard token
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   183
  language of ML is augmented by special syntactic entities of the
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   184
  following form:
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   185
39829
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   186
  \indexouternonterm{antiquote}
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   187
  \begin{rail}
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   188
  antiquote: atsign lbrace nameref args rbrace | lbracesym | rbracesym
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   189
  ;
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   190
  \end{rail}
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   191
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   192
  Here @{syntax nameref} and @{syntax args} are regular outer syntax
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   193
  categories.  Note that attributes and proof methods use similar
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   194
  syntax.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   195
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   196
  \medskip A regular antiquotation @{text "@{name args}"} processes
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   197
  its arguments by the usual means of the Isar source language, and
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   198
  produces corresponding ML source text, either as literal
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   199
  \emph{inline} text (e.g. @{text "@{term t}"}) or abstract
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   200
  \emph{value} (e.g. @{text "@{thm th}"}).  This pre-compilation
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   201
  scheme allows to refer to formal entities in a robust manner, with
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   202
  proper static scoping and with some degree of logical checking of
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   203
  small portions of the code.
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   204
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   205
  Special antiquotations like @{text "@{let \<dots>}"} or @{text "@{note
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   206
  \<dots>}"} augment the compilation context without generating code.  The
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   207
  non-ASCII braces @{text "\<lbrace>"} and @{text "\<rbrace>"} allow to delimit the
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   208
  effect by introducing local blocks within the pre-compilation
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   209
  environment.
39829
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   210
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   211
  \medskip See also \cite{Wenzel-Chaieb:2007b} for a slightly broader
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   212
  perspective on Isabelle/ML antiquotations.  *}
39829
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   213
39830
7c501d7f1e45 clarified tag markup;
wenzelm
parents: 39829
diff changeset
   214
text %mlantiq {*
39829
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   215
  \begin{matharray}{rcl}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   216
  @{ML_antiquotation_def "let"} & : & @{text ML_antiquotation} \\
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   217
  @{ML_antiquotation_def "note"} & : & @{text ML_antiquotation} \\
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   218
  \end{matharray}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   219
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   220
  \begin{rail}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   221
  'let' ((term + 'and') '=' term + 'and')
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   222
  ;
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   223
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   224
  'note' (thmdef? thmrefs + 'and')
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   225
  ;
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   226
  \end{rail}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   227
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   228
  \begin{description}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   229
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   230
  \item @{text "@{let p = t}"} binds schematic variables in the
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   231
  pattern @{text "p"} by higher-order matching against the term @{text
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   232
  "t"}.  This is analogous to the regular @{command_ref let} command
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   233
  in the Isar proof language.  The pre-compilation environment is
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   234
  augmented by auxiliary term bindings, without emitting ML source.
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   235
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   236
  \item @{text "@{note a = b\<^sub>1 \<dots> b\<^sub>n}"} recalls existing facts @{text
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   237
  "b\<^sub>1, \<dots>, b\<^sub>n"}, binding the result as @{text a}.  This is analogous to
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   238
  the regular @{command_ref note} command in the Isar proof language.
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   239
  The pre-compilation environment is augmented by auxiliary fact
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   240
  bindings, without emitting ML source.
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   241
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   242
  \end{description}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   243
*}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   244
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   245
text %mlex {* The following artificial example gives a first
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   246
  impression about using the antiquotation elements introduced so far,
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   247
  together with the basic @{text "@{thm}"} antiquotation defined
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   248
  later.
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   249
*}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   250
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   251
ML {*
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   252
  \<lbrace>
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   253
    @{let ?t = my_term}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   254
    @{note my_refl = reflexive [of ?t]}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   255
    fun foo th = Thm.transitive th @{thm my_refl}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   256
  \<rbrace>
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   257
*}
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   258
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   259
text {*
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   260
  The extra block delimiters do not affect the compiled code itself, i.e.\
f5ea50decd9f more on ML antiquotations;
wenzelm
parents: 39827
diff changeset
   261
  function @{ML foo} is available in the present context.
39827
d829ce302ca4 basic setup for ML antiquotations -- with rail diagrams;
wenzelm
parents: 39825
diff changeset
   262
*}
39823
61e7935a6858 more on Isabelle/ML;
wenzelm
parents: 39822
diff changeset
   263
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   264
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   265
section {* Message output channels \label{sec:message-channels} *}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   266
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   267
text {* Isabelle provides output channels for different kinds of
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   268
  messages: regular output, high-volume tracing information, warnings,
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   269
  and errors.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   270
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   271
  Depending on the user interface involved, these messages may appear
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   272
  in different text styles or colours.  The standard output for
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   273
  terminal sessions prefixes each line of warnings by @{verbatim
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   274
  "###"} and errors by @{verbatim "***"}, but leaves anything else
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   275
  unchanged.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   276
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   277
  Messages are associated with the transaction context of the running
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   278
  Isar command.  This enables the front-end to manage commands and
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   279
  resulting messages together.  For example, after deleting a command
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   280
  from a given theory document version, the corresponding message
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   281
  output can be retracted from the display.  *}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   282
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   283
text %mlref {*
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   284
  \begin{mldecls}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   285
  @{index_ML writeln: "string -> unit"} \\
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   286
  @{index_ML tracing: "string -> unit"} \\
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   287
  @{index_ML warning: "string -> unit"} \\
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   288
  @{index_ML error: "string -> 'a"} \\
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   289
  \end{mldecls}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   290
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   291
  \begin{description}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   292
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   293
  \item @{ML writeln}~@{text "text"} outputs @{text "text"} as regular
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   294
  message.  This is the primary message output operation of Isabelle
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   295
  and should be used by default.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   296
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   297
  \item @{ML tracing}~@{text "text"} outputs @{text "text"} as special
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   298
  tracing message, indicating potential high-volume output to the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   299
  front-end (hundreds or thousands of messages issued by a single
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   300
  command).  The idea is to allow the user-interface to downgrade the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   301
  quality of message display to achieve higher throughput.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   302
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   303
  Note that the user might have to take special actions to see tracing
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   304
  output, e.g.\ switch to a different output window.  So this channel
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   305
  should not be used for regular output.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   306
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   307
  \item @{ML warning}~@{text "text"} outputs @{text "text"} as
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   308
  warning, which typically means some extra emphasis on the front-end
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   309
  side (color highlighting, icon).
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   310
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   311
  \item @{ML error}~@{text "text"} raises exception @{ML ERROR}~@{text
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   312
  "text"} and thus lets the Isar toplevel print @{text "text"} on the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   313
  error channel, which typically means some extra emphasis on the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   314
  front-end side (color highlighting, icon).
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   315
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   316
  This assumes that the exception is not handled before the command
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   317
  terminates.  Handling exception @{ML ERROR}~@{text "text"} is a
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   318
  perfectly legal alternative: it means that the error is absorbed
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   319
  without any message output.
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   320
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   321
  \begin{warn}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   322
  The actual error channel is accessed via @{ML Output.error_msg}, but
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   323
  the interaction protocol of Proof~General \emph{crashes} if that
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   324
  function is used in regular ML code: error output and toplevel
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   325
  command failure always need to coincide here.
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   326
  \end{warn}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   327
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   328
  \end{description}
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   329
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   330
  \begin{warn}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   331
  Regular Isabelle/ML code should output messages exclusively by the
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   332
  official channels.  Using raw I/O on \emph{stdout} or \emph{stderr}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   333
  instead (e.g.\ via @{ML TextIO.output}) is apt to cause problems in
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   334
  the presence of parallel and asynchronous processing of Isabelle
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   335
  theories.  Such raw output might be displayed by the front-end in
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   336
  some system console log, with a low chance that the user will ever
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   337
  see it.  Moreover, as a genuine side-effect on global process
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   338
  channels, there is no proper way to retract output when Isar command
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   339
  transactions are reset.
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39859
diff changeset
   340
  \end{warn}
39835
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   341
*}
6cefd96bb71d modernized version of "Message output channels";
wenzelm
parents: 39830
diff changeset
   342
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   343
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   344
section {* Exceptions \label{sec:exceptions} *}
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   345
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   346
text {* The Standard ML semantics of strict functional evaluation
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   347
  together with exceptions is rather well defined, but some delicate
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   348
  points need to be observed to avoid that ML programs go wrong
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   349
  despite static type-checking.  Exceptions in Isabelle/ML are
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   350
  subsequently categorized as follows.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   351
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   352
  \paragraph{Regular user errors.}  These are meant to provide
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   353
  informative feedback about malformed input etc.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   354
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   355
  The \emph{error} function raises the corresponding \emph{ERROR}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   356
  exception, with a plain text message as argument.  \emph{ERROR}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   357
  exceptions can be handled internally, in order to be ignored, turned
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   358
  into other exceptions, or cascaded by appending messages.  If the
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   359
  corresponding Isabelle/Isar command terminates with an \emph{ERROR}
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   360
  exception state, the toplevel will print the result on the error
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   361
  channel (see \secref{sec:message-channels}).
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   362
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   363
  It is considered bad style to refer to internal function names or
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   364
  values in ML source notation in user error messages.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   365
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   366
  Grammatical correctness of error messages can be improved by
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   367
  \emph{omitting} final punctuation: messages are often concatenated
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   368
  or put into a larger context (e.g.\ augmented with source position).
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   369
  By not insisting in the final word at the origin of the error, the
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   370
  system can perform its administrative tasks more easily and
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   371
  robustly.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   372
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   373
  \paragraph{Program failures.}  There is a handful of standard
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   374
  exceptions that indicate general failure situations, or failures of
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   375
  core operations on logical entities (types, terms, theorems,
39856
wenzelm
parents: 39855
diff changeset
   376
  theories, see \chref{ch:logic}).
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   377
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   378
  These exceptions indicate a genuine breakdown of the program, so the
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   379
  main purpose is to determine quickly what has happened where.
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   380
  Traditionally, the (short) exception message would include the name
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   381
  of an ML function, although this no longer necessary, because the ML
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   382
  runtime system prints a detailed source position of the
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   383
  corresponding @{verbatim raise} keyword.
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   384
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   385
  \medskip User modules can always introduce their own custom
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   386
  exceptions locally, e.g.\ to organize internal failures robustly
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   387
  without overlapping with existing exceptions.  Exceptions that are
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   388
  exposed in module signatures require extra care, though, and should
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   389
  \emph{not} be introduced by default.  Surprise by end-users or ML
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   390
  users of a module can be often minimized by using plain user errors.
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   391
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   392
  \paragraph{Interrupts.}  These indicate arbitrary system events:
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   393
  both the ML runtime system and the Isabelle/ML infrastructure signal
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   394
  various exceptional situations by raising the special
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   395
  \emph{Interrupt} exception in user code.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   396
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   397
  This is the one and only way that physical events can intrude an
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   398
  Isabelle/ML program.  Such an interrupt can mean out-of-memory,
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   399
  stack overflow, timeout, internal signaling of threads, or the user
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   400
  producing a console interrupt manually etc.  An Isabelle/ML program
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   401
  that intercepts interrupts becomes dependent on physical effects of
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   402
  the environment.  Even worse, exception handling patterns that are
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   403
  too general by accident, e.g.\ by mispelled exception constructors,
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   404
  will cover interrupts unintentionally, and thus render the program
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   405
  semantics ill-defined.
39854
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   406
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   407
  Note that the Interrupt exception dates back to the original SML90
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   408
  language definition.  It was excluded from the SML97 version to
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   409
  avoid its malign impact on ML program semantics, but without
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   410
  providing a viable alternative.  Isabelle/ML recovers physical
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   411
  interruptibility (which an indispensable tool to implement managed
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   412
  evaluation of Isar command transactions), but requires user code to
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   413
  be strictly transparent wrt.\ interrupts.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   414
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   415
  \begin{warn}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   416
  Isabelle/ML user code needs to terminate promptly on interruption,
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   417
  without guessing at its meaning to the system infrastructure.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   418
  Temporary handling of interrupts for cleanup of global resources
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   419
  etc.\ needs to be followed immediately by re-raising of the original
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   420
  exception.
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   421
  \end{warn}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   422
*}
7480a7a159cb more on "Exceptions";
wenzelm
parents: 39851
diff changeset
   423
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   424
text %mlref {*
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   425
  \begin{mldecls}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   426
  @{index_ML try: "('a -> 'b) -> 'a -> 'b option"} \\
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   427
  @{index_ML can: "('a -> 'b) -> 'a -> bool"} \\
39856
wenzelm
parents: 39855
diff changeset
   428
  @{index_ML ERROR: "string -> exn"} \\
wenzelm
parents: 39855
diff changeset
   429
  @{index_ML Fail: "string -> exn"} \\
wenzelm
parents: 39855
diff changeset
   430
  @{index_ML Exn.is_interrupt: "exn -> bool"} \\
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   431
  @{index_ML reraise: "exn -> 'a"} \\
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   432
  @{index_ML exception_trace: "(unit -> 'a) -> 'a"} \\
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   433
  \end{mldecls}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   434
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   435
  \begin{description}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   436
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   437
  \item @{ML try}~@{text "f x"} makes the partiality of evaluating
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   438
  @{text "f x"} explicit via the option datatype.  Interrupts are
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   439
  \emph{not} handled here, i.e.\ this form serves as safe replacement
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   440
  for the \emph{unsafe} version @{verbatim "(SOME"}~@{text "f
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   441
  x"}~@{verbatim "handle _ => NONE)"} that is occasionally seen in
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   442
  books about SML.
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   443
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   444
  \item @{ML can} is similar to @{ML try} with more abstract result.
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   445
39856
wenzelm
parents: 39855
diff changeset
   446
  \item @{ML ERROR}~@{text "msg"} represents user errors; this
wenzelm
parents: 39855
diff changeset
   447
  exception is always raised via the @{ML error} function (see
wenzelm
parents: 39855
diff changeset
   448
  \secref{sec:message-channels}).
wenzelm
parents: 39855
diff changeset
   449
wenzelm
parents: 39855
diff changeset
   450
  \item @{ML Fail}~@{text "msg"} represents general program failures.
wenzelm
parents: 39855
diff changeset
   451
wenzelm
parents: 39855
diff changeset
   452
  \item @{ML Exn.is_interrupt} identifies interrupts robustly, without
wenzelm
parents: 39855
diff changeset
   453
  mentioning concrete exception constructors in user code.  Handled
wenzelm
parents: 39855
diff changeset
   454
  interrupts need to be re-raised promptly!
wenzelm
parents: 39855
diff changeset
   455
39855
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   456
  \item @{ML reraise}~@{text "exn"} raises exception @{text "exn"}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   457
  while preserving its implicit position information (if possible,
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   458
  depending on the ML platform).
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   459
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   460
  \item @{ML exception_trace}~@{verbatim "(fn _ =>"}~@{text
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   461
  "e"}@{verbatim ")"} evaluates expression @{text "e"} while printing
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   462
  a full trace of its stack of nested exceptions (if possible,
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   463
  depending on the ML platform).\footnote{In various versions of
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   464
  Poly/ML the trace will appear on raw stdout of the Isabelle
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   465
  process.}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   466
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   467
  Inserting @{ML exception_trace} into ML code temporarily is useful
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   468
  for debugging, but not suitable for production code.
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   469
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   470
  \end{description}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   471
*}
d4299b415879 more on "Exceptions";
wenzelm
parents: 39854
diff changeset
   472
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   473
text %mlantiq {*
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   474
  \begin{matharray}{rcl}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   475
  @{ML_antiquotation_def "assert"} & : & @{text ML_antiquotation} \\
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   476
  \end{matharray}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   477
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   478
  \begin{description}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   479
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   480
  \item @{text "@{assert}"} inlines a function @{text "bool \<Rightarrow> unit"}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   481
  that raises @{ML Fail} if the argument is @{ML false}.  Due to
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   482
  inlining the source position of failed assertions is included in the
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   483
  error output.
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   484
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   485
  \end{description}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   486
*}
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   487
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   488
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   489
section {* Basic data types *}
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   490
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   491
text {* The basis library proposal of SML97 need to be treated with
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   492
  caution.  Many of its operations simply do not fit with important
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   493
  Isabelle/ML conventions (like ``canonical argument order'', see
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   494
  \secref{sec:canonical-argument-order}), others can cause serious
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   495
  problems with the parallel evaluation model of Isabelle/ML (such as
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   496
  @{ML TextIO.print} or @{ML OS.Process.system}).
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   497
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   498
  Subsequently we give a brief overview of important operations on
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   499
  basic ML data types.
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   500
*}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   501
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   502
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   503
subsection {* Characters *}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   504
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   505
text %mlref {*
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   506
  \begin{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   507
  @{index_ML_type char} \\
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   508
  \end{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   509
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   510
  \begin{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   511
39864
wenzelm
parents: 39863
diff changeset
   512
  \item Type @{ML_type char} is \emph{not} used.  The smallest textual
wenzelm
parents: 39863
diff changeset
   513
  unit in Isabelle is represented a ``symbol'' (see
wenzelm
parents: 39863
diff changeset
   514
  \secref{sec:symbols}).
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   515
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   516
  \end{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   517
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   518
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   519
39862
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   520
subsection {* Integers *}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   521
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   522
text %mlref {*
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   523
  \begin{mldecls}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   524
  @{index_ML_type int} \\
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   525
  \end{mldecls}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   526
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   527
  \begin{description}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   528
39864
wenzelm
parents: 39863
diff changeset
   529
  \item Type @{ML_type int} represents regular mathematical integers,
wenzelm
parents: 39863
diff changeset
   530
  which are \emph{unbounded}.  Overflow never happens in
39862
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   531
  practice.\footnote{The size limit for integer bit patterns in memory
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   532
  is 64\,MB for 32-bit Poly/ML, and much higher for 64-bit systems.}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   533
  This works uniformly for all supported ML platforms (Poly/ML and
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   534
  SML/NJ).
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   535
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   536
  Literal integers in ML text (e.g.\ @{ML
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   537
  123456789012345678901234567890}) are forced to be of this one true
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   538
  integer type --- overloading of SML97 is disabled.
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   539
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   540
  Structure @{ML_struct IntInf} of SML97 is obsolete, always use
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   541
  @{ML_struct Int}.  Structure @{ML_struct Integer} in @{"file"
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   542
  "~~/src/Pure/General/integer.ML"} provides some additional
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   543
  operations.
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   544
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   545
  \end{description}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   546
*}
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   547
78e6ec83762a more on "Integers";
wenzelm
parents: 39861
diff changeset
   548
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   549
subsection {* Options *}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   550
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   551
text %mlref {*
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   552
  \begin{mldecls}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   553
  @{index_ML Option.map: "('a -> 'b) -> 'a option -> 'b option"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   554
  @{index_ML is_some: "'a option -> bool"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   555
  @{index_ML is_none: "'a option -> bool"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   556
  @{index_ML the: "'a option -> 'a"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   557
  @{index_ML these: "'a list option -> 'a list"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   558
  @{index_ML the_list: "'a option -> 'a list"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   559
  @{index_ML the_default: "'a -> 'a option -> 'a"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   560
  \end{mldecls}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   561
*}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   562
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   563
text {* Apart from @{ML Option.map} most operations defined in
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   564
  structure @{ML_struct Option} are alien to Isabelle/ML.  The
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   565
  operations shown above are defined in @{"file"
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   566
  "~~/src/Pure/General/basics.ML"}, among others.  *}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   567
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   568
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   569
subsection {* Lists *}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   570
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   571
text {* Lists are ubiquitous in ML as simple and light-weight
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   572
  ``collections'' for many everyday programming tasks.  Isabelle/ML
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   573
  provides some important refinements over the predefined operations
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   574
  in SML97. *}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   575
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   576
text %mlref {*
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   577
  \begin{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   578
  @{index_ML cons: "'a -> 'a list -> 'a list"} \\
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   579
  \end{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   580
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   581
  \begin{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   582
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   583
  \item @{ML cons}~@{text "x xs"} evaluates to @{text "x :: xs"}.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   584
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   585
  Tupled infix operators are a historical accident in Standard ML.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   586
  The curried @{ML cons} amends this, but it should be only used when
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   587
  partial application is required.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   588
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   589
  \end{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   590
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   591
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   592
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   593
subsubsection {* Canonical iteration *}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   594
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   595
text {* A function @{text "f: \<alpha> \<rightarrow> \<beta> \<rightarrow> \<beta>"} can be understood as update
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   596
  on a configuration of type @{text "\<beta>"} that is parametrized by
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   597
  arguments of type @{text "\<alpha>"}.  Given @{text "a: \<alpha>"} the partial
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   598
  application @{text "(f a): \<beta> \<rightarrow> \<beta>"} operates homogeneously on @{text
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   599
  "\<beta>"}.  This can be iterated naturally over a list of parameters
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   600
  @{text "[a\<^sub>1, \<dots>, a\<^sub>n]"} as @{text "f a\<^sub>1; \<dots>; f a\<^bsub>n\<^esub>\<^bsub>\<^esub>"} (where the
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   601
  semicolon represents left-to-right composition).  The latter
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   602
  expression is again a function @{text "\<beta> \<rightarrow> \<beta>"}.  It can be applied
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   603
  to an initial configuration @{text "b: \<beta>"} to start the iteration
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   604
  over the given list of arguments: each @{text "a"} in @{text "a\<^sub>1, \<dots>,
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   605
  a\<^sub>n"} is applied consecutively by updating a cumulative
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   606
  configuration.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   607
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   608
  The @{text fold} combinator in Isabelle/ML lifts a function @{text
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   609
  "f"} as above to its iterated version over a list of arguments.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   610
  Lifting can be repeated, e.g.\ @{text "(fold \<circ> fold) f"} iterates
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   611
  over a list of lists as expected.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   612
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   613
  The variant @{text "fold_rev"} works inside-out over the list of
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   614
  arguments, such that @{text "fold_rev f \<equiv> fold f \<circ> rev"} holds.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   615
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   616
  The @{text "fold_map"} combinator essentially performs @{text
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   617
  "fold"} and @{text "map"} simultaneously: each application of @{text
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   618
  "f"} produces an updated configuration together with a side-result;
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   619
  the iteration collects all such side-results as a separate list.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   620
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   621
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   622
text %mlref {*
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   623
  \begin{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   624
  @{index_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   625
  @{index_ML fold_rev: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   626
  @{index_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   627
  \end{mldecls}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   628
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   629
  \begin{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   630
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   631
  \item @{ML fold}~@{text f} lifts the parametrized update function
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   632
  @{text "f"} to a list of parameters.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   633
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   634
  \item @{ML fold_rev}~@{text "f"} is similar to @{ML fold}~@{text
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   635
  "f"}, but works inside-out.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   636
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   637
  \item @{ML fold_map}~@{text "f"} lifts the parametrized update
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   638
  function @{text "f"} (with side-result) to a list of parameters and
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   639
  cumulative side-results.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   640
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   641
  \end{description}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   642
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   643
  \begin{warn}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   644
  The literature on functional programming provides a multitude of
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   645
  combinators called @{text "foldl"}, @{text "foldr"} etc.  SML97
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   646
  provides its own variations as @{ML List.foldl} and @{ML
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   647
  List.foldr}, while the classic Isabelle library still has the
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   648
  slightly more convenient historic @{ML Library.foldl} and @{ML
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   649
  Library.foldr}.  To avoid further confusion, all of this should be
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   650
  ignored, and @{ML fold} (or @{ML fold_rev}) used exclusively.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   651
  \end{warn}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   652
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   653
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   654
text %mlex {* Using canonical @{ML fold} together with canonical @{ML
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   655
  cons}, or similar standard operations, alternates the orientation of
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   656
  data.  The is quite natural and should not altered forcible by
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   657
  inserting extra applications @{ML rev}.  The alternative @{ML
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   658
  fold_rev} can be used in the relatively few situations, where
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   659
  alternation should be prevented.
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   660
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   661
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   662
ML {*
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   663
  val items = [1, 2, 3, 4, 5, 6, 7, 8, 9, 10];
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   664
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   665
  val list1 = fold cons items [];
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   666
  @{assert} (list1 = rev items);
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   667
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   668
  val list2 = fold_rev cons items [];
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39864
diff changeset
   669
  @{assert} (list2 = items);
39863
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   670
*}
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   671
c0de5386017e more on "Basic data types";
wenzelm
parents: 39862
diff changeset
   672
39859
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   673
subsection {* Unsynchronized references *}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   674
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   675
text %mlref {*
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   676
  \begin{mldecls}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   677
  @{index_ML Unsynchronized.ref: "'a -> 'a Unsynchronized.ref"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   678
  @{index_ML "!": "'a Unsynchronized.ref -> 'a"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   679
  @{index_ML "op :=": "'a Unsynchronized.ref * 'a -> unit"} \\
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   680
  \end{mldecls}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   681
*}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   682
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   683
text {* Due to ubiquitous parallelism in Isabelle/ML (see also
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   684
  \secref{sec:multi-threading}), the mutable reference cells of
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   685
  Standard ML are notorious for causing problems.  In a highly
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   686
  parallel system, both correctness \emph{and} performance are easily
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   687
  degraded when using mutable data.
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   688
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   689
  The unwieldy name of @{ML Unsynchronized.ref} for the constructor
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   690
  for references in Isabelle/ML emphasizes the inconveniences caused by
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   691
  mutability.  Existing operations @{ML "!"}  and @{ML "op :="} are
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   692
  unchanged, but should be used with special precautions, say in a
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   693
  strictly local situation that is guaranteed to be restricted to
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   694
  sequential evaluation -- now and in the future.  *}
381e16bb5f46 more on "Basic ML data types";
wenzelm
parents: 39856
diff changeset
   695
39867
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   696
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   697
section {* Thread-safe programming *}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   698
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   699
text {* Multi-threaded execution has become an everyday reality in
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   700
  Isabelle since Poly/ML 5.2.1 and Isabelle2008.  Isabelle/ML provides
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   701
  implicit and explicit parallelism by default, without any option to
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   702
  ``drop out''.  User-code that is purely functional and does not
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   703
  intercept interrupts (\secref{sec:exceptions}) can participate
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   704
  within the multi-threaded environment without further ado.  More
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   705
  ambitious tools need to observe the principles explained below.  *}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   706
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   707
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   708
subsection {* Multi-threading with shared memory *}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   709
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   710
text {*
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   711
  Multiple threads help to organize advanced operations of the system,
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   712
  such as real-time conditions on command transactions, sub-components
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   713
  with explicit communication, general asynchronous interaction etc.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   714
  Moreover, truely parallel evaluation is inevitable to make adequate
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   715
  use of the CPU resources that are available on multi-core
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   716
  systems.\footnote{Multi-core computing does not mean that there are
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   717
  ``spare cycles'' to be wasted.  It means that the continued
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   718
  exponential speedup of CPU performance due to ``Moore's Law''
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   719
  follows different rules: clock frequency has reached its peak around
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   720
  2005, and applications need to be parallelized in order to avoid a
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   721
  perceived loss of performance, see also \cite{Sutter:2005}.}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   722
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   723
  Isabelle/Isar exploits the inherent structure of theories and proofs
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   724
  to support \emph{implicit parallelism} to a large extent.  LCF-style
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   725
  theorem provides unusually good conditions for parallelism; see also
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   726
  \cite{Wenzel:2009}.  This means, significant parts of theory and
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   727
  proof checking is parallelized by default.  A maximum speedup-factor
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   728
  of 3.0 on 4 cores and 5.0 on 8 cores can be
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   729
  expected.\footnote{Further scalability is limited due to garbage
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   730
  collection, which is still sequential in Poly/ML 5.2/5.3/5.4.  It
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   731
  helps to provide initial heap space generously, using the
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   732
  \texttt{-H} option.  Initial heap size needs to be scaled-up
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   733
  together with the number of CPU cores.}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   734
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   735
  \medskip ML threads lack the memory protection of separate
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   736
  processes, and operate concurrently on shared heap memory.  This has
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   737
  the advantage that results of independent computations are
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   738
  immediately available to other threads.  Abstract values can be
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   739
  passed between threads without copying or awkward serialization that
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   740
  is typically required for explicit message passing between separate
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   741
  processes.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   742
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   743
  To make shared-memory multi-threading work robustly and efficiently,
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   744
  some programming guidelines need to be observed.  While the ML
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   745
  system is responsible to maintain basic integrity of the
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   746
  representation of ML values in memory, the application programmer
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   747
  needs to ensure that multi-threaded execution does not break the
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   748
  intended semantics.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   749
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   750
  \begin{warn}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   751
  To participate in implicit parallelism, tools need to be
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   752
  thread-safe.  A single ill-behaved tool can affect the stability and
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   753
  performance of the whole system.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   754
  \end{warn}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   755
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   756
  Apart from observing the principles of thread-safeness passively,
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   757
  advanced tools may also exploit parallelism actively, e.g.\ by using
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   758
  ``future values'' (\secref{sec:futures}) or or the more basic
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   759
  library functions for parallel list operations
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   760
  (\secref{sec:parlist}).
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   761
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   762
  \begin{warn}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   763
  Parallel computing resources are managed centrally by the
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   764
  Isabelle/ML infrastructure.  User-space programs must not fork their
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   765
  own threads to perform computations.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   766
  \end{warn}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   767
*}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   768
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   769
subsection {* Critical shared resources *}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   770
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   771
text {* Thread-safeness is mainly concerned about concurrent
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   772
  read/write access to shared resources, which are outside the purely
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   773
  functional world of ML.  This covers the following in particular.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   774
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   775
  \begin{itemize}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   776
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   777
  \item Global references (or arrays), i.e.\ mutable memory cells that
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   778
  persist over several invocations of associated
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   779
  operations.\footnote{This is independent of the visibility of such
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   780
  mutable values in the toplevel scope.}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   781
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   782
  \item Global state of the running process, i.e.\ raw I/O channels,
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   783
  environment variables, current working directory.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   784
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   785
  \item Writable resources in the file-system that are shared among
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   786
  different threads or other processes.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   787
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   788
  \end{itemize}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   789
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   790
  Isabelle/ML provides various mechanisms to \emph{avoid} critical
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   791
  shared resources in most practical situations.  As last resort there
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   792
  are some mechanisms for explicit synchronization.  The following
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   793
  guidelines help to make Isabelle/ML programs work smoothly in a
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   794
  highly parallel environment.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   795
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   796
  \begin{itemize}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   797
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   798
  \item Avoid global references altogether.  Isabelle/Isar maintains a
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   799
  uniform context that incorporates arbitrary data declared by
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   800
  Isabelle/ML programs in user-space (see \secref{sec:context-data}).
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   801
  This context is passed as plain value and user tools can get/map
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   802
  their own data in a purely functional manner.  Configuration options
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   803
  within the context (\secref{sec:config-options}) provide simple
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   804
  drop-in replacements for formerly writable flags.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   805
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   806
  \item Keep components with local state information re-entrant.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   807
  Instead of poking initial values into (private) global references,
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   808
  create a new state record on each invocation, and pass that through
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   809
  any auxiliary functions of the component.  The state record may well
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   810
  contain mutable references, without requiring any special
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   811
  synchronizations, as long as each invocation sees its own copy.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   812
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   813
  \item Raw output on @{text "stdout"} or @{text "stderr"} should be
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   814
  avoided altogether, or at least performed as a single atomic
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   815
  action.\footnote{The Poly/ML library is thread-safe for each
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   816
  individual output operation, but the ordering of parallel
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   817
  invocations is arbitrary.  This means raw output will appear on some
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   818
  system console with unpredictable interleaving of atomic chunks.}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   819
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   820
  Note that regular message output channels
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   821
  (\secref{sec:message-channels}) are not directly affected: each
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   822
  message is associated with the command transaction from where it
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   823
  originates, independently of other transactions.  This means each
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   824
  running command has effectively its own set of message channels, and
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   825
  interleaving can only happen (at message boundary) when commands use
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   826
  parallelism internally.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   827
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   828
  \item Environment variables and the current working directory of the
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   829
  running Isabelle process are considered strictly read-only.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   830
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   831
  \item Writable files are in most situations just temporary input to
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   832
  external processes invoked by some ML thread.  By ensuring a unique
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   833
  file name for each instance, such write operations will be disjoint
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   834
  over the life-time of a given Isabelle process, and thus
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   835
  thread-safe.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   836
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   837
  \end{itemize}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   838
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   839
  In rare situations where actual mutable content needs to be
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   840
  manipulated, Isabelle provides a single \emph{critical section} that
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   841
  may be entered while preventing any other thread from doing the
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   842
  same.  Entering the critical section without contention is very
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   843
  fast, and several basic system operations do so frequently.  This
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   844
  also means that each thread should leave the critical section
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   845
  quickly, otherwise parallel execution performance may degrade
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   846
  significantly.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   847
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   848
  Despite this potential bottle-neck, centralized locking is
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   849
  convenient, because it prevents deadlocks without demanding special
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   850
  precautions.  Explicit communication demands other means, though.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   851
  The high-level abstraction of synchronized variables @{"file"
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   852
  "~~/src/Pure/Concurrent/synchronized.ML"} enables parallel
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   853
  components to communicate via shared state; see also @{"file"
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   854
  "~~/src/Pure/Concurrent/mailbox.ML"} as canonical example.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   855
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   856
*}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   857
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   858
text %mlref {*
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   859
  \begin{mldecls}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   860
  @{index_ML NAMED_CRITICAL: "string -> (unit -> 'a) -> 'a"} \\
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   861
  @{index_ML CRITICAL: "(unit -> 'a) -> 'a"} \\
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   862
  \end{mldecls}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   863
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   864
  \begin{description}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   865
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   866
  \item @{ML NAMED_CRITICAL}~@{text "name f"} evaluates @{text "f ()"}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   867
  while staying within the critical section of Isabelle/Isar.  No
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   868
  other thread may do so at the same time, but non-critical parallel
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   869
  execution will continue.  The @{text "name"} argument serves for
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   870
  diagnostic purposes and might help to spot sources of congestion.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   871
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   872
  \item @{ML CRITICAL} is the same as @{ML NAMED_CRITICAL} with empty
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   873
  name argument.
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   874
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   875
  \end{description}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   876
*}
a8363532cd4d somewhat modernized version of "Thread-safe programming";
wenzelm
parents: 39866
diff changeset
   877
30270
61811c9224a6 dummy changes to produce a new changeset of these files;
wenzelm
parents: 29755
diff changeset
   878
end