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