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