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