doc-src/IsarImplementation/Thy/ML_old.thy
author wenzelm
Sun, 17 Oct 2010 20:25:36 +0100
changeset 39861 b8d89db3e238
parent 39859 381e16bb5f46
child 39867 a8363532cd4d
permissions -rw-r--r--
use continental paragraph style, which works better with mixture of (in)formal text; tuned skips and indents; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39822
0de42180febe basic setup for Chapter 0: Isabelle/ML;
wenzelm
parents: 36164
diff changeset
     1
theory ML_old
29755
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
26437
5906619c8c6b tuned appendix;
wenzelm
parents: 25608
diff changeset
     5
chapter {* Advanced ML programming *}
18538
88fe84d4d151 outline;
wenzelm
parents: 18537
diff changeset
     6
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
     7
section {* Style *}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
     8
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
     9
text {*
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    10
  Like any style guide, also this one should not be interpreted dogmatically, but
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    11
  with care and discernment.  It consists of a collection of
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    12
  recommendations which have been turned out useful over long years of
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    13
  Isabelle system development and are supposed to support writing of readable
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    14
  and managable code.  Special cases encourage derivations,
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    15
  as far as you know what you are doing.
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    16
  \footnote{This style guide is loosely based on
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    17
  \url{http://caml.inria.fr/resources/doc/guides/guidelines.en.html}}
21148
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    18
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    19
  \begin{description}
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    20
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    21
    \item[fundamental law of programming]
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    22
      Whenever writing code, keep in mind: A program is
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    23
      written once, modified ten times, and read
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    24
      hundred times.  So simplify its writing,
21148
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    25
      always keep future modifications in mind,
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    26
      and never jeopardize readability.  Every second you hesitate
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    27
      to spend on making your code more clear you will
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    28
      have to spend ten times understanding what you have
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    29
      written later on.
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    30
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    31
    \item[white space matters]
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    32
      Treat white space in your code as if it determines
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    33
      the meaning of code.
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    34
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    35
      \begin{itemize}
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    36
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    37
        \item The space bar is the easiest key to find on the keyboard,
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    38
          press it as often as necessary. @{verbatim "2 + 2"} is better
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    39
          than @{verbatim "2+2"}, likewise @{verbatim "f (x, y)"} is
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    40
          better than @{verbatim "f(x,y)"}.
21148
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    41
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    42
        \item Restrict your lines to 80 characters.  This will allow
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    43
          you to keep the beginning of a line in view while watching
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    44
          its end.\footnote{To acknowledge the lax practice of
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    45
          text editing these days, we tolerate as much as 100
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    46
          characters per line, but anything beyond 120 is not
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    47
          considered proper source text.}
21148
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    48
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    49
        \item Ban tabulators; they are a context-sensitive formatting
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    50
          feature and likely to confuse anyone not using your favorite
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    51
          editor.\footnote{Some modern programming language even
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    52
          forbid tabulators altogether according to the formal syntax
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    53
          definition.}
21148
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    54
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    55
        \item Get rid of trailing whitespace.  Instead, do not
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    56
          suppress a trailing newline at the end of your files.
21148
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    57
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    58
        \item Choose a generally accepted style of indentation,
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    59
          then use it systematically throughout the whole
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    60
          application.  An indentation of two spaces is appropriate.
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    61
          Avoid dangling indentation.
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    62
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    63
      \end{itemize}
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    64
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    65
    \item[cut-and-paste succeeds over copy-and-paste]
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    66
       \emph{Never} copy-and-paste code when programming.  If you
21148
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    67
        need the same piece of code twice, introduce a
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    68
        reasonable auxiliary function (if there is no
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    69
        such function, very likely you got something wrong).
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    70
        Any copy-and-paste will turn out to be painful 
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    71
        when something has to be changed later on.
21148
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    72
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    73
    \item[comments]
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    74
      are a device which requires careful thinking before using
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    75
      it.  The best comment for your code should be the code itself.
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    76
      Prefer efforts to write clear, understandable code
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    77
      over efforts to explain nasty code.
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    78
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    79
    \item[functional programming is based on functions]
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    80
      Model things as abstract as appropriate.  Avoid unnecessarrily
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    81
      concrete datatype  representations.  For example, consider a function
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    82
      taking some table data structure as argument and performing
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    83
      lookups on it.  Instead of taking a table, it could likewise
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
    84
      take just a lookup function.  Accustom
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    85
      your way of coding to the level of expressiveness a functional
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    86
      programming language is giving onto you.
21148
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    87
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    88
    \item[tuples]
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    89
      are often in the way.  When there is no striking argument
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    90
      to tuple function arguments, just write your function curried.
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    91
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
    92
    \item[telling names]
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    93
      Any name should tell its purpose as exactly as possible, while
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    94
      keeping its length to the absolutely necessary minimum.  Always
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    95
      give the same name to function arguments which have the same
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    96
      meaning. Separate words by underscores (@{verbatim
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    97
      int_of_string}, not @{verbatim intOfString}).\footnote{Some
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    98
      recent tools for Emacs include special precautions to cope with
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
    99
      bumpy names in @{text "camelCase"}, e.g.\ for improved on-screen
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   100
      readability.  It is easier to abstain from using such names in the
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   101
      first place.}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   102
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   103
  \end{description}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   104
*}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   105
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   106
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   107
section {* Thread-safe programming *}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   108
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   109
text {*
29159
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   110
  Recent versions of Poly/ML (5.2.1 or later) support robust
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   111
  multithreaded execution, based on native operating system threads of
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   112
  the underlying platform.  Thus threads will actually be executed in
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   113
  parallel on multi-core systems.  A speedup-factor of approximately
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   114
  1.5--3 can be expected on a regular 4-core machine.\footnote{There
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   115
  is some inherent limitation of the speedup factor due to garbage
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   116
  collection, which is still sequential.  It helps to provide initial
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   117
  heap space generously, using the \texttt{-H} option of Poly/ML.}
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   118
  Threads also help to organize advanced operations of the system,
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   119
  with explicit communication between sub-components, real-time
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   120
  conditions, time-outs etc.
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   121
29159
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   122
  Threads lack the memory protection of separate processes, and
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   123
  operate concurrently on shared heap memory.  This has the advantage
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   124
  that results of independent computations are immediately available
29159
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   125
  to other threads, without requiring untyped character streams,
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   126
  awkward serialization etc.
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   127
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   128
  On the other hand, some programming guidelines need to be observed
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   129
  in order to make unprotected parallelism work out smoothly.  While
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   130
  the ML system implementation is responsible to maintain basic
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   131
  integrity of the representation of ML values in memory, the
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   132
  application programmer needs to ensure that multithreaded execution
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   133
  does not break the intended semantics.
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   134
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   135
  \medskip \paragraph{Critical shared resources.} Actually only those
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   136
  parts outside the purely functional world of ML are critical.  In
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   137
  particular, this covers
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   138
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   139
  \begin{itemize}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   140
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   141
  \item global references (or arrays), i.e.\ those that persist over
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   142
  several invocations of associated operations,\footnote{This is
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   143
  independent of the visibility of such mutable values in the toplevel
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   144
  scope.}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   145
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   146
  \item direct I/O on shared channels, notably @{text "stdin"}, @{text
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   147
  "stdout"}, @{text "stderr"}.
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   148
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   149
  \end{itemize}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   150
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   151
  The majority of tools implemented within the Isabelle/Isar framework
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   152
  will not require any of these critical elements: nothing special
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   153
  needs to be observed when staying in the purely functional fragment
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   154
  of ML.  Note that output via the official Isabelle channels does not
29159
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   155
  count as direct I/O, so the operations @{ML "writeln"}, @{ML
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   156
  "warning"}, @{ML "tracing"} etc.\ are safe.
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   157
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   158
  Moreover, ML bindings within the toplevel environment (@{verbatim
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   159
  "type"}, @{verbatim val}, @{verbatim "structure"} etc.) due to
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   160
  run-time invocation of the compiler are also safe, because
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   161
  Isabelle/Isar manages this as part of the theory or proof context.
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   162
29159
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   163
  \paragraph{Multithreading in Isabelle/Isar.}  The theory loader
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   164
  automatically exploits the overall parallelism of independent nodes
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   165
  in the development graph, as well as the inherent irrelevance of
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   166
  proofs for goals being fully specified in advance.  This means,
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   167
  checking of individual Isar proofs is parallelized by default.
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   168
  Beyond that, very sophisticated proof tools may use local
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   169
  parallelism internally, via the general programming model of
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   170
  ``future values'' (see also @{"file"
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   171
  "~~/src/Pure/Concurrent/future.ML"}).
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   172
29159
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   173
  Any ML code that works relatively to the present background theory
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   174
  is already safe.  Contextual data may be easily stored within the
24090
wenzelm
parents: 24089
diff changeset
   175
  theory or proof context, thanks to the generic data concept of
wenzelm
parents: 24089
diff changeset
   176
  Isabelle/Isar (see \secref{sec:context-data}).  This greatly
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   177
  diminishes the demand for global state information in the first
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   178
  place.
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   179
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   180
  \medskip In rare situations where actual mutable content needs to be
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   181
  manipulated, Isabelle provides a single \emph{critical section} that
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   182
  may be entered while preventing any other thread from doing the
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   183
  same.  Entering the critical section without contention is very
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   184
  fast, and several basic system operations do so frequently.  This
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   185
  also means that each thread should leave the critical section
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   186
  quickly, otherwise parallel execution performance may degrade
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   187
  significantly.
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   188
29159
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   189
  Despite this potential bottle-neck, centralized locking is
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   190
  convenient, because it prevents deadlocks without demanding special
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   191
  precautions.  Explicit communication demands other means, though.
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   192
  The high-level abstraction of synchronized variables @{"file"
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   193
  "~~/src/Pure/Concurrent/synchronized.ML"} enables parallel
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   194
  components to communicate via shared state; see also @{"file"
e876b8bc0f77 updated thread-safe programming;
wenzelm
parents: 29008
diff changeset
   195
  "~~/src/Pure/Concurrent/mailbox.ML"} as canonical example.
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   196
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   197
  \paragraph{Good conduct of impure programs.} The following
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   198
  guidelines enable non-functional programs to participate in
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   199
  multithreading.
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   200
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   201
  \begin{itemize}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   202
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   203
  \item Minimize global state information.  Using proper theory and
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   204
  proof context data will actually return to functional update of
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   205
  values, without any special precautions for multithreading.  Apart
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   206
  from the fully general functors for theory and proof data (see
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   207
  \secref{sec:context-data}) there are drop-in replacements that
24090
wenzelm
parents: 24089
diff changeset
   208
  emulate primitive references for common cases of \emph{configuration
wenzelm
parents: 24089
diff changeset
   209
  options} for type @{ML_type "bool"}/@{ML_type "int"}/@{ML_type
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24090
diff changeset
   210
  "string"} (see structure @{ML_struct Config} and @{ML
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24090
diff changeset
   211
  Attrib.config_bool} etc.), and lists of theorems (see functor
36164
532f4d1cb0fc salvaged some ML functors from decay, which is the natural consequence of lack of formal checking;
wenzelm
parents: 36134
diff changeset
   212
  @{ML_functor Named_Thms}).
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   213
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   214
  \item Keep components with local state information
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   215
  \emph{re-entrant}.  Instead of poking initial values into (private)
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   216
  global references, create a new state record on each invocation, and
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   217
  pass that through any auxiliary functions of the component.  The
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   218
  state record may well contain mutable references, without requiring
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   219
  any special synchronizations, as long as each invocation sees its
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   220
  own copy.  Occasionally, one might even return to plain functional
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   221
  updates on non-mutable record values here.
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   222
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   223
  \item Isolate process configuration flags.  The main legitimate
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   224
  application of global references is to configure the whole process
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   225
  in a certain way, essentially affecting all threads.  A typical
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   226
  example is the @{ML show_types} flag, which tells the pretty printer
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   227
  to output explicit type information for terms.  Such flags usually
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   228
  do not affect the functionality of the core system, but only the
24090
wenzelm
parents: 24089
diff changeset
   229
  view being presented to the user.
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   230
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   231
  Occasionally, such global process flags are treated like implicit
32966
5b21661fe618 indicate CRITICAL nature of various setmp combinators;
wenzelm
parents: 32833
diff changeset
   232
  arguments to certain operations, by using the @{ML setmp_CRITICAL} combinator
24090
wenzelm
parents: 24089
diff changeset
   233
  for safe temporary assignment.  Its traditional purpose was to
wenzelm
parents: 24089
diff changeset
   234
  ensure proper recovery of the original value when exceptions are
wenzelm
parents: 24089
diff changeset
   235
  raised in the body, now the functionality is extended to enter the
wenzelm
parents: 24089
diff changeset
   236
  \emph{critical section} (with its usual potential of degrading
wenzelm
parents: 24089
diff changeset
   237
  parallelism).
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   238
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   239
  Note that recovery of plain value passing semantics via @{ML
32966
5b21661fe618 indicate CRITICAL nature of various setmp combinators;
wenzelm
parents: 32833
diff changeset
   240
  setmp_CRITICAL}~@{text "ref value"} assumes that this @{text "ref"} is
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   241
  exclusively manipulated within the critical section.  In particular,
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   242
  any persistent global assignment of @{text "ref := value"} needs to
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   243
  be marked critical as well, to prevent intruding another threads
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   244
  local view, and a lost-update in the global scope, too.
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   245
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   246
  \end{itemize}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   247
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   248
  Recall that in an open ``LCF-style'' system like Isabelle/Isar, the
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   249
  user participates in constructing the overall environment.  This
24090
wenzelm
parents: 24089
diff changeset
   250
  means that state-based facilities offered by one component will
wenzelm
parents: 24089
diff changeset
   251
  require special caution later on.  So minimizing critical elements,
wenzelm
parents: 24089
diff changeset
   252
  by staying within the plain value-oriented view relative to theory
wenzelm
parents: 24089
diff changeset
   253
  or proof contexts most of the time, will also reduce the chance of
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   254
  mishaps occurring to end-users.
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   255
*}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   256
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   257
text %mlref {*
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   258
  \begin{mldecls}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   259
  @{index_ML NAMED_CRITICAL: "string -> (unit -> 'a) -> 'a"} \\
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   260
  @{index_ML CRITICAL: "(unit -> 'a) -> 'a"} \\
32966
5b21661fe618 indicate CRITICAL nature of various setmp combinators;
wenzelm
parents: 32833
diff changeset
   261
  @{index_ML setmp_CRITICAL: "'a Unsynchronized.ref -> 'a -> ('b -> 'c) -> 'b -> 'c"} \\
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   262
  \end{mldecls}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   263
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   264
  \begin{description}
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   265
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   266
  \item @{ML NAMED_CRITICAL}~@{text "name f"} evaluates @{text "f ()"}
24090
wenzelm
parents: 24089
diff changeset
   267
  while staying within the critical section of Isabelle/Isar.  No
wenzelm
parents: 24089
diff changeset
   268
  other thread may do so at the same time, but non-critical parallel
wenzelm
parents: 24089
diff changeset
   269
  execution will continue.  The @{text "name"} argument serves for
wenzelm
parents: 24089
diff changeset
   270
  diagnostic purposes and might help to spot sources of congestion.
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   271
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   272
  \item @{ML CRITICAL} is the same as @{ML NAMED_CRITICAL} with empty
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   273
  name argument.
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   274
32966
5b21661fe618 indicate CRITICAL nature of various setmp combinators;
wenzelm
parents: 32833
diff changeset
   275
  \item @{ML setmp_CRITICAL}~@{text "ref value f x"} evaluates @{text "f x"}
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   276
  while staying within the critical section and having @{text "ref :=
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   277
  value"} assigned temporarily.  This recovers a value-passing
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   278
  semantics involving global references, regardless of exceptions or
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   279
  concurrency.
21148
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
   280
3a64d58a9f49 first version of style guide
haftmann
parents: 20520
diff changeset
   281
  \end{description}
18554
bff7a1466fe4 more stuff;
wenzelm
parents: 18538
diff changeset
   282
*}
18538
88fe84d4d151 outline;
wenzelm
parents: 18537
diff changeset
   283
88fe84d4d151 outline;
wenzelm
parents: 18537
diff changeset
   284
88fe84d4d151 outline;
wenzelm
parents: 18537
diff changeset
   285
chapter {* Basic library functions *}
88fe84d4d151 outline;
wenzelm
parents: 18537
diff changeset
   286
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   287
text {*
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   288
  Beyond the proposal of the SML/NJ basis library, Isabelle comes
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   289
  with its own library, from which selected parts are given here.
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   290
  These should encourage a study of the Isabelle sources,
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   291
  in particular files \emph{Pure/library.ML} and \emph{Pure/General/*.ML}.
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   292
*}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   293
29646
5941c156902d added label;
wenzelm
parents: 29612
diff changeset
   294
section {* Linear transformations \label{sec:ML-linear-trans} *}
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   295
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   296
text %mlref {*
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   297
  \begin{mldecls}
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   298
  @{index_ML "op |> ": "'a * ('a -> 'b) -> 'b"} \\
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   299
  \end{mldecls}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   300
*}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   301
22322
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   302
(*<*)
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   303
typedecl foo
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   304
consts foo :: foo
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   305
ML {*
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   306
val thy = Theory.copy @{theory}
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   307
*}
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   308
(*>*)
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   309
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   310
text {*
26460
21504de31197 some styling
haftmann
parents: 26459
diff changeset
   311
  \noindent Many problems in functional programming can be thought of
22322
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   312
  as linear transformations, i.e.~a caluclation starts with a
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   313
  particular value @{ML_text "x : foo"} which is then transformed
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   314
  by application of a function @{ML_text "f : foo -> foo"},
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   315
  continued by an application of a function @{ML_text "g : foo -> bar"},
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   316
  and so on.  As a canoncial example, take functions enriching
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   317
  a theory by constant declararion and primitive definitions:
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   318
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   319
  \smallskip\begin{mldecls}
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 32966
diff changeset
   320
  @{ML "Sign.declare_const: (binding * typ) * mixfix
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   321
  -> theory -> term * theory"} \\
36134
c210a8fda4c5 updated Thm.add_axiom/add_def;
wenzelm
parents: 33174
diff changeset
   322
  @{ML "Thm.add_def: bool -> bool -> binding * term -> theory -> (string * thm) * theory"}
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   323
  \end{mldecls}
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   324
26460
21504de31197 some styling
haftmann
parents: 26459
diff changeset
   325
  \noindent Written with naive application, an addition of constant
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   326
  @{term bar} with type @{typ "foo \<Rightarrow> foo"} and
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   327
  a corresponding definition @{term "bar \<equiv> \<lambda>x. x"} would look like:
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   328
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   329
  \smallskip\begin{mldecls}
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   330
  @{ML "(fn (t, thy) => Thm.add_def false false
29612
4f68e0f8f4fd binding replaces bstring
haftmann
parents: 29581
diff changeset
   331
  (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})) thy)
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 32966
diff changeset
   332
    (Sign.declare_const
29008
d863c103ded0 adapted to changes in binding module
haftmann
parents: 28283
diff changeset
   333
      ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn) thy)"}
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   334
  \end{mldecls}
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   335
26460
21504de31197 some styling
haftmann
parents: 26459
diff changeset
   336
  \noindent With increasing numbers of applications, this code gets quite
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   337
  unreadable.  Further, it is unintuitive that things are
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   338
  written down in the opposite order as they actually ``happen''.
22322
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   339
*}
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   340
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   341
text {*
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   342
  \noindent At this stage, Isabelle offers some combinators which allow
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   343
  for more convenient notation, most notably reverse application:
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   344
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   345
  \smallskip\begin{mldecls}
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   346
@{ML "thy
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 32966
diff changeset
   347
|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   348
|> (fn (t, thy) => thy
25608
7793d6423d01 adjusted
haftmann
parents: 25185
diff changeset
   349
|> Thm.add_def false false
29612
4f68e0f8f4fd binding replaces bstring
haftmann
parents: 29581
diff changeset
   350
     (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))"}
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   351
  \end{mldecls}
22322
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   352
*}
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   353
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   354
text %mlref {*
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   355
  \begin{mldecls}
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   356
  @{index_ML "op |-> ": "('c * 'a) * ('c -> 'a -> 'b) -> 'b"} \\
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   357
  @{index_ML "op |>> ": "('a * 'c) * ('a -> 'b) -> 'b * 'c"} \\
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   358
  @{index_ML "op ||> ": "('c * 'a) * ('a -> 'b) -> 'c * 'b"} \\
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   359
  @{index_ML "op ||>> ": "('c * 'a) * ('a -> 'd * 'b) -> ('c * 'd) * 'b"} \\
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   360
  \end{mldecls}
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   361
*}
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   362
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   363
text {*
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   364
  \noindent Usually, functions involving theory updates also return
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   365
  side results; to use these conveniently, yet another
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   366
  set of combinators is at hand, most notably @{ML "op |->"}
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   367
  which allows curried access to side results:
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   368
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   369
  \smallskip\begin{mldecls}
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   370
@{ML "thy
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 32966
diff changeset
   371
|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
25608
7793d6423d01 adjusted
haftmann
parents: 25185
diff changeset
   372
|-> (fn t => Thm.add_def false false
29612
4f68e0f8f4fd binding replaces bstring
haftmann
parents: 29581
diff changeset
   373
      (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   374
"}
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   375
  \end{mldecls}
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   376
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   377
  \noindent @{ML "op |>>"} allows for processing side results on their own:
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   378
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   379
  \smallskip\begin{mldecls}
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   380
@{ML "thy
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 32966
diff changeset
   381
|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   382
|>> (fn t => Logic.mk_equals (t, @{term \"%x. x\"}))
29612
4f68e0f8f4fd binding replaces bstring
haftmann
parents: 29581
diff changeset
   383
|-> (fn def => Thm.add_def false false (Binding.name \"bar_def\", def))
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   384
"}
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   385
  \end{mldecls}
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   386
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   387
  \noindent Orthogonally, @{ML "op ||>"} applies a transformation
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   388
  in the presence of side results which are left unchanged:
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   389
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   390
  \smallskip\begin{mldecls}
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   391
@{ML "thy
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 32966
diff changeset
   392
|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   393
||> Sign.add_path \"foobar\"
25608
7793d6423d01 adjusted
haftmann
parents: 25185
diff changeset
   394
|-> (fn t => Thm.add_def false false
29612
4f68e0f8f4fd binding replaces bstring
haftmann
parents: 29581
diff changeset
   395
      (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   396
||> Sign.restore_naming thy
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   397
"}
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   398
  \end{mldecls}
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   399
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   400
  \noindent @{ML "op ||>>"} accumulates side results:
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   401
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   402
  \smallskip\begin{mldecls}
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   403
@{ML "thy
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 32966
diff changeset
   404
|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 32966
diff changeset
   405
||>> Sign.declare_const ((Binding.name \"foobar\", @{typ \"foo => foo\"}), NoSyn)
25608
7793d6423d01 adjusted
haftmann
parents: 25185
diff changeset
   406
|-> (fn (t1, t2) => Thm.add_def false false
29612
4f68e0f8f4fd binding replaces bstring
haftmann
parents: 29581
diff changeset
   407
      (Binding.name \"bar_def\", Logic.mk_equals (t1, t2)))
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   408
"}
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   409
  \end{mldecls}
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   410
*}
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   411
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   412
text %mlref {*
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   413
  \begin{mldecls}
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   414
  @{index_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   415
  @{index_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   416
  \end{mldecls}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   417
*}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   418
22322
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   419
text {*
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   420
  \noindent This principles naturally lift to \emph{lists} using
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   421
  the @{ML fold} and @{ML fold_map} combinators.
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   422
  The first lifts a single function
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   423
  \begin{quote}\footnotesize
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   424
    @{ML_text "f : 'a -> 'b -> 'b"} to @{ML_text "'a list -> 'b -> 'b"}
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   425
  \end{quote}
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   426
  such that
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   427
  \begin{quote}\footnotesize
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   428
    @{ML_text "y |> fold f [x1, x2, ..., x_n]"} \\
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   429
    \hspace*{2ex}@{text "\<leadsto>"} @{ML_text "y |> f x1 |> f x2 |> ... |> f x_n"}
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   430
  \end{quote}
26460
21504de31197 some styling
haftmann
parents: 26459
diff changeset
   431
  \noindent The second accumulates side results in a list by lifting
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   432
  a single function
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   433
  \begin{quote}\footnotesize
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   434
    @{ML_text "f : 'a -> 'b -> 'c * 'b"} to @{ML_text "'a list -> 'b -> 'c list * 'b"}
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   435
  \end{quote}
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   436
  such that
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   437
  \begin{quote}\footnotesize
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   438
    @{ML_text "y |> fold_map f [x1, x2, ..., x_n]"} \\
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   439
    \hspace*{2ex}@{text "\<leadsto>"} @{ML_text "y |> f x1 ||>> f x2 ||>> ... ||>> f x_n"} \\
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   440
    \hspace*{6ex}@{ML_text "||> (fn ((z1, z2), ..., z_n) => [z1, z2, ..., z_n])"}
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   441
  \end{quote}
25185
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   442
  
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   443
  \noindent Example:
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   444
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   445
  \smallskip\begin{mldecls}
25185
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   446
@{ML "let
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   447
  val consts = [\"foo\", \"bar\"];
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   448
in
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   449
  thy
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 32966
diff changeset
   450
  |> fold_map (fn const => Sign.declare_const
29008
d863c103ded0 adapted to changes in binding module
haftmann
parents: 28283
diff changeset
   451
       ((Binding.name const, @{typ \"foo => foo\"}), NoSyn)) consts
25185
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   452
  |>> map (fn t => Logic.mk_equals (t, @{term \"%x. x\"}))
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   453
  |-> (fn defs => fold_map (fn def =>
29612
4f68e0f8f4fd binding replaces bstring
haftmann
parents: 29581
diff changeset
   454
       Thm.add_def false false (Binding.empty, def)) defs)
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   455
end"}
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   456
  \end{mldecls}
22322
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   457
*}
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   458
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   459
text %mlref {*
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   460
  \begin{mldecls}
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   461
  @{index_ML "op #> ": "('a -> 'b) * ('b -> 'c) -> 'a -> 'c"} \\
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   462
  @{index_ML "op #-> ": "('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd"} \\
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   463
  @{index_ML "op #>> ": "('a -> 'c * 'b) * ('c -> 'd) -> 'a -> 'd * 'b"} \\
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   464
  @{index_ML "op ##> ": "('a -> 'c * 'b) * ('b -> 'd) -> 'a -> 'c * 'd"} \\
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   465
  @{index_ML "op ##>> ": "('a -> 'c * 'b) * ('b -> 'e * 'd) -> 'a -> ('c * 'e) * 'd"} \\
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   466
  \end{mldecls}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   467
*}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   468
22322
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   469
text {*
22550
c5039bee2602 updated
haftmann
parents: 22503
diff changeset
   470
  \noindent All those linear combinators also exist in higher-order
c5039bee2602 updated
haftmann
parents: 22503
diff changeset
   471
  variants which do not expect a value on the left hand side
c5039bee2602 updated
haftmann
parents: 22503
diff changeset
   472
  but a function.
22322
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   473
*}
b9924abb8c66 continued
haftmann
parents: 22293
diff changeset
   474
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   475
text %mlref {*
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   476
  \begin{mldecls}
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   477
  @{index_ML "op ` ": "('b -> 'a) -> 'b -> 'a * 'b"} \\
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   478
  @{index_ML tap: "('b -> 'a) -> 'b -> 'b"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   479
  \end{mldecls}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   480
*}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   481
22550
c5039bee2602 updated
haftmann
parents: 22503
diff changeset
   482
text {*
25185
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   483
  \noindent These operators allow to ``query'' a context
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   484
  in a series of context transformations:
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   485
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   486
  \smallskip\begin{mldecls}
25185
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   487
@{ML "thy
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   488
|> tap (fn _ => writeln \"now adding constant\")
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 32966
diff changeset
   489
|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
25185
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   490
||>> `(fn thy => Sign.declared_const thy
29008
d863c103ded0 adapted to changes in binding module
haftmann
parents: 28283
diff changeset
   491
         (Sign.full_name thy (Binding.name \"foobar\")))
25185
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   492
|-> (fn (t, b) => if b then I
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 32966
diff changeset
   493
       else Sign.declare_const
29008
d863c103ded0 adapted to changes in binding module
haftmann
parents: 28283
diff changeset
   494
         ((Binding.name \"foobar\", @{typ \"foo => foo\"}), NoSyn) #> snd)
25185
762ed22d15c7 continued
haftmann
parents: 25151
diff changeset
   495
"}
26459
bb0e729be5a4 some styling
haftmann
parents: 26437
diff changeset
   496
  \end{mldecls}
22550
c5039bee2602 updated
haftmann
parents: 22503
diff changeset
   497
*}
c5039bee2602 updated
haftmann
parents: 22503
diff changeset
   498
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   499
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   500
section {* Common data structures *}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   501
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   502
subsection {* Lists (as set-like data structures) *}
18538
88fe84d4d151 outline;
wenzelm
parents: 18537
diff changeset
   503
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   504
text {*
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   505
  \begin{mldecls}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   506
  @{index_ML member: "('b * 'a -> bool) -> 'a list -> 'b -> bool"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   507
  @{index_ML insert: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   508
  @{index_ML remove: "('b * 'a -> bool) -> 'b -> 'a list -> 'a list"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   509
  @{index_ML merge: "('a * 'a -> bool) -> 'a list * 'a list -> 'a list"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   510
  \end{mldecls}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   511
*}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   512
22503
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   513
text {*
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   514
  Lists are often used as set-like data structures -- set-like in
25151
9374a0df240c continued
haftmann
parents: 24110
diff changeset
   515
  the sense that they support a notion of @{ML member}-ship,
22503
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   516
  @{ML insert}-ing and @{ML remove}-ing, but are order-sensitive.
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   517
  This is convenient when implementing a history-like mechanism:
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   518
  @{ML insert} adds an element \emph{to the front} of a list,
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   519
  if not yet present; @{ML remove} removes \emph{all} occurences
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   520
  of a particular element.  Correspondingly @{ML merge} implements a 
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   521
  a merge on two lists suitable for merges of context data
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   522
  (\secref{sec:context-theory}).
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   523
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   524
  Functions are parametrized by an explicit equality function
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   525
  to accomplish overloaded equality;  in most cases of monomorphic
24089
070d539ba403 tuned section "Style";
wenzelm
parents: 23652
diff changeset
   526
  equality, writing @{ML "op ="} should suffice.
22503
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   527
*}
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   528
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   529
subsection {* Association lists *}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   530
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   531
text {*
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   532
  \begin{mldecls}
23652
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   533
  @{index_ML_exn AList.DUP} \\
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   534
  @{index_ML AList.lookup: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   535
  @{index_ML AList.defined: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   536
  @{index_ML AList.update: "('a * 'a -> bool) -> ('a * 'b) -> ('a * 'b) list -> ('a * 'b) list"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   537
  @{index_ML AList.default: "('a * 'a -> bool) -> ('a * 'b) -> ('a * 'b) list -> ('a * 'b) list"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   538
  @{index_ML AList.delete: "('a * 'b -> bool) -> 'a -> ('b * 'c) list -> ('b * 'c) list"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   539
  @{index_ML AList.map_entry: "('a * 'b -> bool) -> 'a
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   540
    -> ('c -> 'c) -> ('b * 'c) list -> ('b * 'c) list"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   541
  @{index_ML AList.map_default: "('a * 'a -> bool) -> 'a * 'b -> ('b -> 'b)
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   542
    -> ('a * 'b) list -> ('a * 'b) list"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   543
  @{index_ML AList.join: "('a * 'a -> bool) -> ('a -> 'b * 'b -> 'b) (*exception DUP*)
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   544
    -> ('a * 'b) list * ('a * 'b) list -> ('a * 'b) list (*exception AList.DUP*)"} \\
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   545
  @{index_ML AList.merge: "('a * 'a -> bool) -> ('b * 'b -> bool)
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   546
    -> ('a * 'b) list * ('a * 'b) list -> ('a * 'b) list (*exception AList.DUP*)"}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   547
  \end{mldecls}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   548
*}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   549
22503
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   550
text {*
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   551
  Association lists can be seens as an extension of set-like lists:
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   552
  on the one hand, they may be used to implement finite mappings,
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   553
  on the other hand, they remain order-sensitive and allow for
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   554
  multiple key-value-pair with the same key: @{ML AList.lookup}
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   555
  returns the \emph{first} value corresponding to a particular
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   556
  key, if present.  @{ML AList.update} updates
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   557
  the \emph{first} occurence of a particular key; if no such
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   558
  key exists yet, the key-value-pair is added \emph{to the front}.
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   559
  @{ML AList.delete} only deletes the \emph{first} occurence of a key.
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   560
  @{ML AList.merge} provides an operation suitable for merges of context data
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   561
  (\secref{sec:context-theory}), where an equality parameter on
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   562
  values determines whether a merge should be considered a conflict.
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   563
  A slightly generalized operation if implementend by the @{ML AList.join}
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   564
  function which allows for explicit conflict resolution.
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   565
*}
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   566
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   567
subsection {* Tables *}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   568
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   569
text {*
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   570
  \begin{mldecls}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   571
  @{index_ML_type "'a Symtab.table"} \\
23652
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   572
  @{index_ML_exn Symtab.DUP: string} \\
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   573
  @{index_ML_exn Symtab.SAME} \\
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   574
  @{index_ML_exn Symtab.UNDEF: string} \\
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   575
  @{index_ML Symtab.empty: "'a Symtab.table"} \\
23652
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   576
  @{index_ML Symtab.lookup: "'a Symtab.table -> string -> 'a option"} \\
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   577
  @{index_ML Symtab.defined: "'a Symtab.table -> string -> bool"} \\
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   578
  @{index_ML Symtab.update: "(string * 'a) -> 'a Symtab.table -> 'a Symtab.table"} \\
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   579
  @{index_ML Symtab.default: "string * 'a -> 'a Symtab.table -> 'a Symtab.table"} \\
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   580
  @{index_ML Symtab.delete: "string
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   581
    -> 'a Symtab.table -> 'a Symtab.table (*exception Symtab.UNDEF*)"} \\
23652
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   582
  @{index_ML Symtab.map_entry: "string -> ('a -> 'a)
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   583
    -> 'a Symtab.table -> 'a Symtab.table"} \\
23652
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   584
  @{index_ML Symtab.map_default: "(string * 'a) -> ('a -> 'a)
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   585
    -> 'a Symtab.table -> 'a Symtab.table"} \\
23652
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   586
  @{index_ML Symtab.join: "(string -> 'a * 'a -> 'a) (*exception Symtab.DUP/Symtab.SAME*)
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   587
    -> 'a Symtab.table * 'a Symtab.table
23652
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   588
    -> 'a Symtab.table (*exception Symtab.DUP*)"} \\
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   589
  @{index_ML Symtab.merge: "('a * 'a -> bool)
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   590
    -> 'a Symtab.table * 'a Symtab.table
23652
94eeb79be496 simplified Symtab;
wenzelm
parents: 22550
diff changeset
   591
    -> 'a Symtab.table (*exception Symtab.DUP*)"}
22293
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   592
  \end{mldecls}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   593
*}
3593a76c9ed3 added outline for Isabelle library description
haftmann
parents: 21502
diff changeset
   594
22503
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   595
text {*
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   596
  Tables are an efficient representation of finite mappings without
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   597
  any notion of order;  due to their efficiency they should be used
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   598
  whenever such pure finite mappings are neccessary.
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   599
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   600
  The key type of tables must be given explicitly by instantiating
36164
532f4d1cb0fc salvaged some ML functors from decay, which is the natural consequence of lack of formal checking;
wenzelm
parents: 36134
diff changeset
   601
  the @{ML_functor Table} functor which takes the key type
22503
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   602
  together with its @{ML_type order}; for convience, we restrict
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   603
  here to the @{ML_struct Symtab} instance with @{ML_type string}
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   604
  as key type.
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   605
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   606
  Most table functions correspond to those of association lists.
d53664118418 added some sketches about library functions
haftmann
parents: 22322
diff changeset
   607
*}
20489
a684fc70d04e tentative appendix C;
wenzelm
parents: 18554
diff changeset
   608
30270
61811c9224a6 dummy changes to produce a new changeset of these files;
wenzelm
parents: 29755
diff changeset
   609
end