18538
|
1 |
(* $Id$ *)
|
|
2 |
|
|
3 |
theory "ML" imports base begin
|
|
4 |
|
|
5 |
chapter {* Aesthetics of ML programming *}
|
|
6 |
|
24089
|
7 |
section {* Style *}
|
|
8 |
|
21502
|
9 |
text FIXME
|
|
10 |
|
21148
|
11 |
text {* This style guide is loosely based on
|
|
12 |
\url{http://caml.inria.fr/resources/doc/guides/guidelines.en.html}.
|
|
13 |
% FIMXE \url{http://www.cs.cornell.edu/Courses/cs312/2003sp/handouts/style.htm}
|
|
14 |
|
24089
|
15 |
Like any style guide, it should not be interpreted dogmatically, but
|
|
16 |
with care and discernment. Instead, it forms a collection of
|
|
17 |
recommendations which, if obeyed, result in code that is not
|
|
18 |
considered to be obfuscated. In certain cases, derivations are
|
|
19 |
encouraged, as far as you know what you are doing.
|
21148
|
20 |
|
|
21 |
\begin{description}
|
|
22 |
|
|
23 |
\item[fundamental law of programming]
|
|
24 |
Whenever writing code, keep in mind: A program is
|
|
25 |
written once, modified ten times, and read
|
|
26 |
100 times. So simplify its writing,
|
|
27 |
always keep future modifications in mind,
|
|
28 |
and never jeopardize readability. Every second you hesitate
|
|
29 |
to spend on making your code more clear you will
|
|
30 |
have to spend ten times understanding what you have
|
|
31 |
written later on.
|
|
32 |
|
|
33 |
\item[white space matters]
|
|
34 |
Treat white space in your code as if it determines
|
|
35 |
the meaning of code.
|
|
36 |
|
|
37 |
\begin{itemize}
|
|
38 |
|
|
39 |
\item The space bar is the easiest key to find on the keyboard,
|
24089
|
40 |
press it as often as necessary. @{verbatim "2 + 2"} is better
|
|
41 |
than @{verbatim "2+2"}, likewise @{verbatim "f (x, y)"} is
|
|
42 |
better than @{verbatim "f(x,y)"}.
|
21148
|
43 |
|
24089
|
44 |
\item Restrict your lines to 80 characters. This will allow
|
|
45 |
you to keep the beginning of a line in view while watching
|
|
46 |
its end.\footnote{To acknowledge the lax practice of
|
|
47 |
text editing these days, we tolerate as much as 100
|
|
48 |
characters per line, but anything beyond 120 is not
|
|
49 |
considered proper source text.}
|
21148
|
50 |
|
24089
|
51 |
\item Ban tabulators; they are a context-sensitive formatting
|
|
52 |
feature and likely to confuse anyone not using your favorite
|
|
53 |
editor.\footnote{Some modern programming language even
|
|
54 |
forbid tabulators altogether according to the formal syntax
|
|
55 |
definition.}
|
21148
|
56 |
|
|
57 |
\item Get rid of trailing whitespace. Instead, do not
|
24089
|
58 |
suppress a trailing newline at the end of your files.
|
21148
|
59 |
|
|
60 |
\item Choose a generally accepted style of indentation,
|
|
61 |
then use it systematically throughout the whole
|
|
62 |
application. An indentation of two spaces is appropriate.
|
|
63 |
Avoid dangling indentation.
|
|
64 |
|
|
65 |
\end{itemize}
|
|
66 |
|
|
67 |
\item[cut-and-paste succeeds over copy-and-paste]
|
24089
|
68 |
\emph{Never} copy-and-paste code when programming. If you
|
21148
|
69 |
need the same piece of code twice, introduce a
|
|
70 |
reasonable auxiliary function (if there is no
|
|
71 |
such function, very likely you got something wrong).
|
|
72 |
Any copy-and-paste will turn out to be painful
|
|
73 |
when something has to be changed or fixed later on.
|
|
74 |
|
|
75 |
\item[comments]
|
|
76 |
are a device which requires careful thinking before using
|
|
77 |
it. The best comment for your code should be the code itself.
|
|
78 |
Prefer efforts to write clear, understandable code
|
|
79 |
over efforts to explain nasty code.
|
|
80 |
|
|
81 |
\item[functional programming is based on functions]
|
24089
|
82 |
Avoid ``constructivisms'', i.e.\ unnecessary concrete datatype
|
|
83 |
representations. Instead model things as abstract as
|
|
84 |
appropriate. For example, pass a table lookup function rather
|
|
85 |
than a concrete table with lookup performed in body. Accustom
|
|
86 |
your way of coding to the level of expressiveness a functional
|
|
87 |
programming language is giving onto you.
|
21148
|
88 |
|
|
89 |
\item[tuples]
|
|
90 |
are often in the way. When there is no striking argument
|
|
91 |
to tuple function arguments, just write your function curried.
|
|
92 |
|
|
93 |
\item[telling names]
|
24089
|
94 |
Any name should tell its purpose as exactly as possible, while
|
|
95 |
keeping its length to the absolutely necessary minimum. Always
|
|
96 |
give the same name to function arguments which have the same
|
|
97 |
meaning. Separate words by underscores (@{verbatim
|
|
98 |
int_of_string}, not @{verbatim intOfString}).\footnote{Some
|
|
99 |
recent tools for Emacs include special precautions to cope with
|
|
100 |
bumpy names in @{text "camelCase"}, e.g.\ for improved on-screen
|
|
101 |
readability. It is easier to abstain from using such names in the
|
|
102 |
first place.}
|
|
103 |
|
|
104 |
\end{description}
|
|
105 |
*}
|
|
106 |
|
|
107 |
|
|
108 |
section {* Thread-safe programming *}
|
|
109 |
|
|
110 |
text {*
|
|
111 |
Recent versions of Poly/ML (5.1 or later) support multithreaded
|
|
112 |
execution based on native operating system threads of the
|
|
113 |
underlying platform. Thus threads will actually be executed in
|
|
114 |
parallel on multi-core systems. A speedup-factor of approximately
|
|
115 |
2--4 can be expected for large well-structured Isabelle sessions,
|
|
116 |
where theories are organized as a graph with sufficiently many
|
|
117 |
independent nodes.
|
|
118 |
|
|
119 |
Threads lack the memory protection of separate processes, but
|
|
120 |
operate concurrently on shared heap memory. This has the advantage
|
|
121 |
that results of independent computations are immediately available
|
|
122 |
to other threads, without requiring explicit communication,
|
|
123 |
reloading, or even recoding of data.
|
|
124 |
|
|
125 |
On the other hand, some programming guidelines need to be observed
|
|
126 |
in order to make unprotected parallelism work out smoothly. While
|
|
127 |
the ML system implementation is responsible to maintain basic
|
|
128 |
integrity of the representation of ML values in memory, the
|
|
129 |
application programmer needs to ensure that multithreaded execution
|
|
130 |
does not break the intended semantics.
|
|
131 |
|
|
132 |
\medskip \paragraph{Critical shared resources.} Actually only those
|
|
133 |
parts outside the purely functional world of ML are critical. In
|
|
134 |
particular, this covers
|
|
135 |
|
|
136 |
\begin{itemize}
|
|
137 |
|
|
138 |
\item global references (or arrays), i.e.\ those that persist over
|
|
139 |
several invocations of associated operations,\footnote{This is
|
|
140 |
independent of the visibility of such mutable values in the toplevel
|
|
141 |
scope.}
|
|
142 |
|
|
143 |
\item global ML bindings in the toplevel environment (@{verbatim
|
|
144 |
"type"}, @{verbatim val}, @{verbatim "structure"} etc.) due to
|
|
145 |
run-time invocation of the compiler,
|
|
146 |
|
|
147 |
\item direct I/O on shared channels, notably @{text "stdin"}, @{text
|
|
148 |
"stdout"}, @{text "stderr"}.
|
|
149 |
|
|
150 |
\end{itemize}
|
|
151 |
|
|
152 |
The majority of tools implemented within the Isabelle/Isar framework
|
|
153 |
will not require any of these critical elements: nothing special
|
|
154 |
needs to be observed when staying in the purely functional fragment
|
|
155 |
of ML. Note that output via the official Isabelle channels does not
|
|
156 |
even count as direct I/O in the above sense, so the operations @{ML
|
|
157 |
"writeln"}, @{ML "warning"}, @{ML "tracing"} etc.\ are safe.
|
|
158 |
|
|
159 |
\paragraph{Multithreading in Isabelle/Isar.} Our parallel execution
|
|
160 |
model is centered around the theory loader. Whenever a given
|
|
161 |
subgraph of theories needs to be updated, the system schedules a
|
|
162 |
number of threads to process the sources as required, while
|
|
163 |
observing their dependencies. Thus concurrency is limited to
|
|
164 |
independent nodes according to the theory import relation.
|
|
165 |
|
|
166 |
Any user-code that works relatively to the present background theory
|
|
167 |
is already safe. Contextual data may be easily stored within the
|
24090
|
168 |
theory or proof context, thanks to the generic data concept of
|
|
169 |
Isabelle/Isar (see \secref{sec:context-data}). This greatly
|
24089
|
170 |
diminishes the demand for global state information in the first
|
|
171 |
place.
|
|
172 |
|
|
173 |
\medskip In rare situations where actual mutable content needs to be
|
|
174 |
manipulated, Isabelle provides a single \emph{critical section} that
|
|
175 |
may be entered while preventing any other thread from doing the
|
|
176 |
same. Entering the critical section without contention is very
|
|
177 |
fast, and several basic system operations do so frequently. This
|
|
178 |
also means that each thread should leave the critical section
|
|
179 |
quickly, otherwise parallel execution performance may degrade
|
|
180 |
significantly.
|
|
181 |
|
|
182 |
Despite this potential bottle-neck, we refrain from fine-grained
|
|
183 |
locking mechanisms: the restriction to a single lock prevents
|
|
184 |
deadlocks without demanding further considerations in user programs.
|
|
185 |
|
|
186 |
\paragraph{Good conduct of impure programs.} The following
|
|
187 |
guidelines enable non-functional programs to participate in
|
|
188 |
multithreading.
|
|
189 |
|
|
190 |
\begin{itemize}
|
|
191 |
|
|
192 |
\item Minimize global state information. Using proper theory and
|
|
193 |
proof context data will actually return to functional update of
|
|
194 |
values, without any special precautions for multithreading. Apart
|
|
195 |
from the fully general functors for theory and proof data (see
|
|
196 |
\secref{sec:context-data}) there are drop-in replacements that
|
24090
|
197 |
emulate primitive references for common cases of \emph{configuration
|
|
198 |
options} for type @{ML_type "bool"}/@{ML_type "int"}/@{ML_type
|
|
199 |
"string"} (see structure @{ML_struct ConfigOption}) and lists of
|
|
200 |
theorems (see functor @{ML_functor NamedThmsFun}).
|
24089
|
201 |
|
|
202 |
\item Keep components with local state information
|
|
203 |
\emph{re-entrant}. Instead of poking initial values into (private)
|
|
204 |
global references, create a new state record on each invocation, and
|
|
205 |
pass that through any auxiliary functions of the component. The
|
|
206 |
state record may well contain mutable references, without requiring
|
|
207 |
any special synchronizations, as long as each invocation sees its
|
|
208 |
own copy. Occasionally, one might even return to plain functional
|
|
209 |
updates on non-mutable record values here.
|
|
210 |
|
|
211 |
\item Isolate process configuration flags. The main legitimate
|
|
212 |
application of global references is to configure the whole process
|
|
213 |
in a certain way, essentially affecting all threads. A typical
|
|
214 |
example is the @{ML show_types} flag, which tells the pretty printer
|
|
215 |
to output explicit type information for terms. Such flags usually
|
|
216 |
do not affect the functionality of the core system, but only the
|
24090
|
217 |
view being presented to the user.
|
24089
|
218 |
|
|
219 |
Occasionally, such global process flags are treated like implicit
|
|
220 |
arguments to certain operations, by using the @{ML setmp} combinator
|
24090
|
221 |
for safe temporary assignment. Its traditional purpose was to
|
|
222 |
ensure proper recovery of the original value when exceptions are
|
|
223 |
raised in the body, now the functionality is extended to enter the
|
|
224 |
\emph{critical section} (with its usual potential of degrading
|
|
225 |
parallelism).
|
24089
|
226 |
|
|
227 |
Note that recovery of plain value passing semantics via @{ML
|
|
228 |
setmp}~@{text "ref value"} assumes that this @{text "ref"} is
|
|
229 |
exclusively manipulated within the critical section. In particular,
|
|
230 |
any persistent global assignment of @{text "ref := value"} needs to
|
|
231 |
be marked critical as well, to prevent intruding another threads
|
|
232 |
local view, and a lost-update in the global scope, too.
|
|
233 |
|
|
234 |
\item Minimize global ML bindings. Processing theories occasionally
|
|
235 |
affects the global ML environment as well. While each ML
|
|
236 |
compilation unit is safe, the order of scheduling of independent
|
|
237 |
declarations might cause problems when composing several modules
|
|
238 |
later on, due to hiding of previous ML names.
|
|
239 |
|
|
240 |
This cannot be helped in general, because the ML toplevel lacks the
|
|
241 |
graph structure of the Isabelle theory space. Nevertheless, some
|
|
242 |
sound conventions of keeping global ML names essentially disjoint
|
|
243 |
(e.g.\ with the help of ML structures) prevents the problem to occur
|
|
244 |
in most practical situations.
|
|
245 |
|
|
246 |
\end{itemize}
|
|
247 |
|
|
248 |
Recall that in an open ``LCF-style'' system like Isabelle/Isar, the
|
|
249 |
user participates in constructing the overall environment. This
|
24090
|
250 |
means that state-based facilities offered by one component will
|
|
251 |
require special caution later on. So minimizing critical elements,
|
|
252 |
by staying within the plain value-oriented view relative to theory
|
|
253 |
or proof contexts most of the time, will also reduce the chance of
|
24089
|
254 |
mishaps occurring to end-users.
|
|
255 |
*}
|
|
256 |
|
|
257 |
text %mlref {*
|
|
258 |
\begin{mldecls}
|
|
259 |
@{index_ML NAMED_CRITICAL: "string -> (unit -> 'a) -> 'a"} \\
|
|
260 |
@{index_ML CRITICAL: "(unit -> 'a) -> 'a"} \\
|
|
261 |
@{index_ML setmp: "'a ref -> 'a -> ('b -> 'c) -> 'b -> 'c"} \\
|
|
262 |
\end{mldecls}
|
|
263 |
|
|
264 |
\begin{description}
|
|
265 |
|
|
266 |
\item @{ML NAMED_CRITICAL}~@{text "name f"} evaluates @{text "f ()"}
|
24090
|
267 |
while staying within the critical section of Isabelle/Isar. No
|
|
268 |
other thread may do so at the same time, but non-critical parallel
|
|
269 |
execution will continue. The @{text "name"} argument serves for
|
|
270 |
diagnostic purposes and might help to spot sources of congestion.
|
24089
|
271 |
|
|
272 |
\item @{ML CRITICAL} is the same as @{ML NAMED_CRITICAL} with empty
|
|
273 |
name argument.
|
|
274 |
|
|
275 |
\item @{ML setmp}~@{text "ref value f x"} evaluates @{text "f x"}
|
|
276 |
while staying within the critical section and having @{text "ref :=
|
|
277 |
value"} assigned temporarily. This recovers a value-passing
|
|
278 |
semantics involving global references, regardless of exceptions or
|
|
279 |
concurrency.
|
21148
|
280 |
|
|
281 |
\end{description}
|
18554
|
282 |
*}
|
18538
|
283 |
|
|
284 |
|
|
285 |
chapter {* Basic library functions *}
|
|
286 |
|
22293
|
287 |
text {*
|
|
288 |
Beyond the proposal of the SML/NJ basis library, Isabelle comes
|
|
289 |
with its own library, from which selected parts are given here.
|
|
290 |
See further files \emph{Pure/library.ML} and \emph{Pure/General/*.ML}.
|
|
291 |
*}
|
|
292 |
|
|
293 |
section {* Linear transformations *}
|
|
294 |
|
|
295 |
text %mlref {*
|
|
296 |
\begin{mldecls}
|
24089
|
297 |
@{index_ML "op |> ": "'a * ('a -> 'b) -> 'b"} \\
|
22293
|
298 |
@{index_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
|
|
299 |
@{index_ML fold_rev: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
|
|
300 |
\end{mldecls}
|
|
301 |
*}
|
|
302 |
|
22322
|
303 |
(*<*)
|
|
304 |
typedecl foo
|
|
305 |
consts foo :: foo
|
|
306 |
ML {*
|
|
307 |
val dummy_const = ("bar", @{typ foo}, NoSyn)
|
|
308 |
val dummy_def = ("bar", @{term foo})
|
|
309 |
val thy = Theory.copy @{theory}
|
|
310 |
*}
|
|
311 |
(*>*)
|
|
312 |
|
|
313 |
text {*
|
|
314 |
Many problems in functional programming can be thought of
|
|
315 |
as linear transformations, i.e.~a caluclation starts with a
|
|
316 |
particular value @{text "x \<Colon> foo"} which is then transformed
|
|
317 |
by application of a function @{text "f \<Colon> foo \<Rightarrow> foo"},
|
|
318 |
continued by an application of a function @{text "g \<Colon> foo \<Rightarrow> bar"},
|
|
319 |
and so on. As a canoncial example, take primitive functions enriching
|
|
320 |
theories by constants and definitions:
|
22550
|
321 |
@{ML "Sign.add_consts_i: (string * typ * mixfix) list -> theory
|
|
322 |
-> theory"}
|
|
323 |
and @{ML "Theory.add_defs_i: bool -> bool
|
|
324 |
-> (bstring * term) list -> theory -> theory"}.
|
22322
|
325 |
Written with naive application, an addition of a constant with
|
|
326 |
a corresponding definition would look like:
|
22550
|
327 |
@{ML "Theory.add_defs_i false false [dummy_def]
|
|
328 |
(Sign.add_consts_i [dummy_const] thy)"}.
|
22322
|
329 |
With increasing numbers of applications, this code gets quite unreadable.
|
|
330 |
Using composition, at least the nesting of brackets may be reduced:
|
22550
|
331 |
@{ML "(Theory.add_defs_i false false [dummy_def] o Sign.add_consts_i
|
|
332 |
[dummy_const]) thy"}.
|
22322
|
333 |
What remains unsatisfactory is that things are written down in the opposite order
|
|
334 |
as they actually ``happen''.
|
|
335 |
*}
|
|
336 |
|
|
337 |
(*<*)
|
|
338 |
ML {*
|
|
339 |
val thy = Theory.copy @{theory}
|
|
340 |
*}
|
|
341 |
(*>*)
|
|
342 |
|
|
343 |
text {*
|
|
344 |
At this stage, Isabelle offers some combinators which allow for more convenient
|
|
345 |
notation, most notably reverse application:
|
|
346 |
@{ML "
|
|
347 |
thy
|
|
348 |
|> Sign.add_consts_i [dummy_const]
|
|
349 |
|> Theory.add_defs_i false false [dummy_def]"}
|
|
350 |
*}
|
|
351 |
|
|
352 |
text {*
|
22550
|
353 |
\noindent When iterating over a list of parameters @{text "[x\<^isub>1, x\<^isub>2, \<dots> x\<^isub>n] \<Colon> 'a list"},
|
22322
|
354 |
the @{ML fold} combinator lifts a single function @{text "f \<Colon> 'a -> 'b -> 'b"}:
|
|
355 |
@{text "y |> fold f [x\<^isub>1, x\<^isub>2, \<dots> x\<^isub>n] \<equiv> y |> f x\<^isub>1 |> f x\<^isub>2 |> \<dots> |> f x\<^isub>n"}
|
|
356 |
*}
|
22293
|
357 |
|
|
358 |
text %mlref {*
|
|
359 |
\begin{mldecls}
|
24089
|
360 |
@{index_ML "op |-> ": "('c * 'a) * ('c -> 'a -> 'b) -> 'b"} \\
|
|
361 |
@{index_ML "op |>> ": "('a * 'c) * ('a -> 'b) -> 'b * 'c"} \\
|
|
362 |
@{index_ML "op ||> ": "('c * 'a) * ('a -> 'b) -> 'c * 'b"} \\
|
|
363 |
@{index_ML "op ||>> ": "('c * 'a) * ('a -> 'd * 'b) -> ('c * 'd) * 'b"} \\
|
22293
|
364 |
@{index_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\
|
|
365 |
\end{mldecls}
|
|
366 |
*}
|
|
367 |
|
22322
|
368 |
text {*
|
22550
|
369 |
\noindent FIXME transformations involving side results
|
22322
|
370 |
*}
|
|
371 |
|
22293
|
372 |
text %mlref {*
|
|
373 |
\begin{mldecls}
|
24089
|
374 |
@{index_ML "op #> ": "('a -> 'b) * ('b -> 'c) -> 'a -> 'c"} \\
|
|
375 |
@{index_ML "op #-> ": "('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd"} \\
|
|
376 |
@{index_ML "op #>> ": "('a -> 'c * 'b) * ('c -> 'd) -> 'a -> 'd * 'b"} \\
|
|
377 |
@{index_ML "op ##> ": "('a -> 'c * 'b) * ('b -> 'd) -> 'a -> 'c * 'd"} \\
|
|
378 |
@{index_ML "op ##>> ": "('a -> 'c * 'b) * ('b -> 'e * 'd) -> 'a -> ('c * 'e) * 'd"} \\
|
22293
|
379 |
\end{mldecls}
|
|
380 |
*}
|
|
381 |
|
22322
|
382 |
text {*
|
22550
|
383 |
\noindent All those linear combinators also exist in higher-order
|
|
384 |
variants which do not expect a value on the left hand side
|
|
385 |
but a function.
|
22322
|
386 |
*}
|
|
387 |
|
22293
|
388 |
text %mlref {*
|
|
389 |
\begin{mldecls}
|
24089
|
390 |
@{index_ML "op ` ": "('b -> 'a) -> 'b -> 'a * 'b"} \\
|
22293
|
391 |
@{index_ML tap: "('b -> 'a) -> 'b -> 'b"} \\
|
|
392 |
\end{mldecls}
|
|
393 |
*}
|
|
394 |
|
22550
|
395 |
text {*
|
|
396 |
\noindent FIXME
|
|
397 |
*}
|
|
398 |
|
22293
|
399 |
section {* Options and partiality *}
|
|
400 |
|
|
401 |
text %mlref {*
|
|
402 |
\begin{mldecls}
|
|
403 |
@{index_ML is_some: "'a option -> bool"} \\
|
|
404 |
@{index_ML is_none: "'a option -> bool"} \\
|
|
405 |
@{index_ML the: "'a option -> 'a"} \\
|
|
406 |
@{index_ML these: "'a list option -> 'a list"} \\
|
|
407 |
@{index_ML the_list: "'a option -> 'a list"} \\
|
|
408 |
@{index_ML the_default: "'a -> 'a option -> 'a"} \\
|
|
409 |
@{index_ML try: "('a -> 'b) -> 'a -> 'b option"} \\
|
|
410 |
@{index_ML can: "('a -> 'b) -> 'a -> bool"} \\
|
|
411 |
\end{mldecls}
|
|
412 |
*}
|
|
413 |
|
22550
|
414 |
text {*
|
24089
|
415 |
Standard selector functions on @{text option}s are provided. The
|
|
416 |
@{ML try} and @{ML can} functions provide a convenient interface for
|
|
417 |
handling exceptions -- both take as arguments a function @{text f}
|
|
418 |
together with a parameter @{text x} and handle any exception during
|
|
419 |
the evaluation of the application of @{text f} to @{text x}, either
|
|
420 |
return a lifted result (@{ML NONE} on failure) or a boolean value
|
|
421 |
(@{ML false} on failure).
|
22550
|
422 |
*}
|
22293
|
423 |
|
24089
|
424 |
|
22293
|
425 |
section {* Common data structures *}
|
|
426 |
|
|
427 |
subsection {* Lists (as set-like data structures) *}
|
18538
|
428 |
|
22293
|
429 |
text {*
|
|
430 |
\begin{mldecls}
|
|
431 |
@{index_ML member: "('b * 'a -> bool) -> 'a list -> 'b -> bool"} \\
|
|
432 |
@{index_ML insert: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
|
|
433 |
@{index_ML remove: "('b * 'a -> bool) -> 'b -> 'a list -> 'a list"} \\
|
|
434 |
@{index_ML merge: "('a * 'a -> bool) -> 'a list * 'a list -> 'a list"} \\
|
|
435 |
\end{mldecls}
|
|
436 |
*}
|
|
437 |
|
22503
|
438 |
text {*
|
|
439 |
Lists are often used as set-like data structures -- set-like in
|
|
440 |
then sense that they support notion of @{ML member}-ship,
|
|
441 |
@{ML insert}-ing and @{ML remove}-ing, but are order-sensitive.
|
|
442 |
This is convenient when implementing a history-like mechanism:
|
|
443 |
@{ML insert} adds an element \emph{to the front} of a list,
|
|
444 |
if not yet present; @{ML remove} removes \emph{all} occurences
|
|
445 |
of a particular element. Correspondingly @{ML merge} implements a
|
|
446 |
a merge on two lists suitable for merges of context data
|
|
447 |
(\secref{sec:context-theory}).
|
|
448 |
|
|
449 |
Functions are parametrized by an explicit equality function
|
|
450 |
to accomplish overloaded equality; in most cases of monomorphic
|
24089
|
451 |
equality, writing @{ML "op ="} should suffice.
|
22503
|
452 |
*}
|
22293
|
453 |
|
|
454 |
subsection {* Association lists *}
|
|
455 |
|
|
456 |
text {*
|
|
457 |
\begin{mldecls}
|
23652
|
458 |
@{index_ML_exn AList.DUP} \\
|
22293
|
459 |
@{index_ML AList.lookup: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option"} \\
|
|
460 |
@{index_ML AList.defined: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool"} \\
|
|
461 |
@{index_ML AList.update: "('a * 'a -> bool) -> ('a * 'b) -> ('a * 'b) list -> ('a * 'b) list"} \\
|
|
462 |
@{index_ML AList.default: "('a * 'a -> bool) -> ('a * 'b) -> ('a * 'b) list -> ('a * 'b) list"} \\
|
|
463 |
@{index_ML AList.delete: "('a * 'b -> bool) -> 'a -> ('b * 'c) list -> ('b * 'c) list"} \\
|
|
464 |
@{index_ML AList.map_entry: "('a * 'b -> bool) -> 'a
|
|
465 |
-> ('c -> 'c) -> ('b * 'c) list -> ('b * 'c) list"} \\
|
|
466 |
@{index_ML AList.map_default: "('a * 'a -> bool) -> 'a * 'b -> ('b -> 'b)
|
|
467 |
-> ('a * 'b) list -> ('a * 'b) list"} \\
|
|
468 |
@{index_ML AList.join: "('a * 'a -> bool) -> ('a -> 'b * 'b -> 'b) (*exception DUP*)
|
|
469 |
-> ('a * 'b) list * ('a * 'b) list -> ('a * 'b) list (*exception AList.DUP*)"} \\
|
|
470 |
@{index_ML AList.merge: "('a * 'a -> bool) -> ('b * 'b -> bool)
|
|
471 |
-> ('a * 'b) list * ('a * 'b) list -> ('a * 'b) list (*exception AList.DUP*)"}
|
|
472 |
\end{mldecls}
|
|
473 |
*}
|
|
474 |
|
22503
|
475 |
text {*
|
|
476 |
Association lists can be seens as an extension of set-like lists:
|
|
477 |
on the one hand, they may be used to implement finite mappings,
|
|
478 |
on the other hand, they remain order-sensitive and allow for
|
|
479 |
multiple key-value-pair with the same key: @{ML AList.lookup}
|
|
480 |
returns the \emph{first} value corresponding to a particular
|
|
481 |
key, if present. @{ML AList.update} updates
|
|
482 |
the \emph{first} occurence of a particular key; if no such
|
|
483 |
key exists yet, the key-value-pair is added \emph{to the front}.
|
|
484 |
@{ML AList.delete} only deletes the \emph{first} occurence of a key.
|
|
485 |
@{ML AList.merge} provides an operation suitable for merges of context data
|
|
486 |
(\secref{sec:context-theory}), where an equality parameter on
|
|
487 |
values determines whether a merge should be considered a conflict.
|
|
488 |
A slightly generalized operation if implementend by the @{ML AList.join}
|
|
489 |
function which allows for explicit conflict resolution.
|
|
490 |
*}
|
22293
|
491 |
|
|
492 |
subsection {* Tables *}
|
|
493 |
|
|
494 |
text {*
|
|
495 |
\begin{mldecls}
|
|
496 |
@{index_ML_type "'a Symtab.table"} \\
|
23652
|
497 |
@{index_ML_exn Symtab.DUP: string} \\
|
|
498 |
@{index_ML_exn Symtab.SAME} \\
|
|
499 |
@{index_ML_exn Symtab.UNDEF: string} \\
|
22293
|
500 |
@{index_ML Symtab.empty: "'a Symtab.table"} \\
|
23652
|
501 |
@{index_ML Symtab.lookup: "'a Symtab.table -> string -> 'a option"} \\
|
|
502 |
@{index_ML Symtab.defined: "'a Symtab.table -> string -> bool"} \\
|
|
503 |
@{index_ML Symtab.update: "(string * 'a) -> 'a Symtab.table -> 'a Symtab.table"} \\
|
|
504 |
@{index_ML Symtab.default: "string * 'a -> 'a Symtab.table -> 'a Symtab.table"} \\
|
|
505 |
@{index_ML Symtab.delete: "string
|
22293
|
506 |
-> 'a Symtab.table -> 'a Symtab.table (*exception Symtab.UNDEF*)"} \\
|
23652
|
507 |
@{index_ML Symtab.map_entry: "string -> ('a -> 'a)
|
22293
|
508 |
-> 'a Symtab.table -> 'a Symtab.table"} \\
|
23652
|
509 |
@{index_ML Symtab.map_default: "(string * 'a) -> ('a -> 'a)
|
22293
|
510 |
-> 'a Symtab.table -> 'a Symtab.table"} \\
|
23652
|
511 |
@{index_ML Symtab.join: "(string -> 'a * 'a -> 'a) (*exception Symtab.DUP/Symtab.SAME*)
|
22293
|
512 |
-> 'a Symtab.table * 'a Symtab.table
|
23652
|
513 |
-> 'a Symtab.table (*exception Symtab.DUP*)"} \\
|
22293
|
514 |
@{index_ML Symtab.merge: "('a * 'a -> bool)
|
|
515 |
-> 'a Symtab.table * 'a Symtab.table
|
23652
|
516 |
-> 'a Symtab.table (*exception Symtab.DUP*)"}
|
22293
|
517 |
\end{mldecls}
|
|
518 |
*}
|
|
519 |
|
22503
|
520 |
text {*
|
|
521 |
Tables are an efficient representation of finite mappings without
|
|
522 |
any notion of order; due to their efficiency they should be used
|
|
523 |
whenever such pure finite mappings are neccessary.
|
|
524 |
|
|
525 |
The key type of tables must be given explicitly by instantiating
|
|
526 |
the @{ML_functor TableFun} functor which takes the key type
|
|
527 |
together with its @{ML_type order}; for convience, we restrict
|
|
528 |
here to the @{ML_struct Symtab} instance with @{ML_type string}
|
|
529 |
as key type.
|
|
530 |
|
|
531 |
Most table functions correspond to those of association lists.
|
|
532 |
*}
|
20489
|
533 |
|
23652
|
534 |
|
20489
|
535 |
chapter {* Cookbook *}
|
|
536 |
|
20491
|
537 |
section {* A method that depends on declarations in the context *}
|
20489
|
538 |
|
|
539 |
text FIXME
|
|
540 |
|
18538
|
541 |
end
|