| author | wenzelm | 
| Sun, 24 Aug 2025 20:26:02 +0200 | |
| changeset 83053 | c1ccd17fb70f | 
| parent 78742 | b2216709a839 | 
| 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: 
39824diff
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: 
39824diff
changeset | 624 |   \end{mldecls}
 | 
| 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 wenzelm parents: 
39824diff
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: 
39824diff
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: 
39824diff
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: 
39825diff
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: 
39825diff
changeset | 659 | |
| 62969 | 660 |   Here @{syntax name} and @{syntax args} are outer syntax categories, as
 | 
| 76987 | 661 | defined in \<^cite>\<open>"isabelle-isar-ref"\<close>. | 
| 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: 
51295diff
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: 
51295diff
changeset | 692 |   \end{matharray}
 | 
| 
e49bf0be79ba
document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
 wenzelm parents: 
51295diff
changeset | 693 | |
| 69597 | 694 | \<^rail>\<open> | 
| 51636 
e49bf0be79ba
document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
 wenzelm parents: 
51295diff
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: 
51295diff
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: 
51295diff
changeset | 716 | val x = 42; | 
| 
e49bf0be79ba
document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
 wenzelm parents: 
51295diff
changeset | 717 | val y = true; | 
| 
e49bf0be79ba
document @{make_string}, cf. NEWS of Isabelle2009-2 (June 2010);
 wenzelm parents: 
51295diff
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 | |
| 76987 | 969 | formatting, which is routinely used in the Prover IDE \<^cite>\<open>"isabelle-jedit"\<close>. | 
| 39835 | 970 | |
| 61854 | 971 | Messages are associated with the transaction context of the running Isar | 
| 972 | command. This enables the front-end to manage commands and resulting | |
| 973 | messages together. For example, after deleting a command from a given theory | |
| 974 | document version, the corresponding message output can be retracted from the | |
| 975 | display. | |
| 58618 | 976 | \<close> | 
| 977 | ||
| 978 | text %mlref \<open> | |
| 39835 | 979 |   \begin{mldecls}
 | 
| 73764 | 980 |   @{define_ML writeln: "string -> unit"} \\
 | 
| 981 |   @{define_ML tracing: "string -> unit"} \\
 | |
| 982 |   @{define_ML warning: "string -> unit"} \\
 | |
| 983 |   @{define_ML error: "string -> 'a"} % FIXME Output.error_message (!?) \\
 | |
| 39835 | 984 |   \end{mldecls}
 | 
| 985 | ||
| 69597 | 986 | \<^descr> \<^ML>\<open>writeln\<close>~\<open>text\<close> outputs \<open>text\<close> as regular message. This is the | 
| 61854 | 987 | primary message output operation of Isabelle and should be used by default. | 
| 39835 | 988 | |
| 69597 | 989 | \<^descr> \<^ML>\<open>tracing\<close>~\<open>text\<close> outputs \<open>text\<close> as special tracing message, indicating | 
| 61854 | 990 | potential high-volume output to the front-end (hundreds or thousands of | 
| 991 | messages issued by a single command). The idea is to allow the | |
| 992 | user-interface to downgrade the quality of message display to achieve higher | |
| 993 | throughput. | |
| 39835 | 994 | |
| 61854 | 995 | Note that the user might have to take special actions to see tracing output, | 
| 996 | e.g.\ switch to a different output window. So this channel should not be | |
| 997 | used for regular output. | |
| 39835 | 998 | |
| 69597 | 999 | \<^descr> \<^ML>\<open>warning\<close>~\<open>text\<close> outputs \<open>text\<close> as warning, which typically means some | 
| 61854 | 1000 | extra emphasis on the front-end side (color highlighting, icons, etc.). | 
| 39835 | 1001 | |
| 69597 | 1002 | \<^descr> \<^ML>\<open>error\<close>~\<open>text\<close> raises exception \<^ML>\<open>ERROR\<close>~\<open>text\<close> and thus lets the | 
| 61854 | 1003 | Isar toplevel print \<open>text\<close> on the error channel, which typically means some | 
| 1004 | extra emphasis on the front-end side (color highlighting, icons, etc.). | |
| 39835 | 1005 | |
| 1006 | This assumes that the exception is not handled before the command | |
| 69597 | 1007 | terminates. Handling exception \<^ML>\<open>ERROR\<close>~\<open>text\<close> is a perfectly legal | 
| 61854 | 1008 | alternative: it means that the error is absorbed without any message output. | 
| 39835 | 1009 | |
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39859diff
changeset | 1010 |   \begin{warn}
 | 
| 69597 | 1011 | The actual error channel is accessed via \<^ML>\<open>Output.error_message\<close>, but | 
| 58842 | 1012 | 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: 
39859diff
changeset | 1013 |   \end{warn}
 | 
| 39835 | 1014 | |
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39859diff
changeset | 1015 | |
| 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39859diff
changeset | 1016 |   \begin{warn}
 | 
| 61854 | 1017 | Regular Isabelle/ML code should output messages exclusively by the official | 
| 75617 | 1018 | channels. Using raw I/O on \<^emph>\<open>stdout\<close> or \<^emph>\<open>stderr\<close> instead (e.g.\ via | 
| 1019 | \<^ML>\<open>TextIO.output\<close>) is apt to cause problems in the presence of parallel | |
| 1020 | and asynchronous processing of Isabelle theories. Such raw output might be | |
| 61854 | 1021 | displayed by the front-end in some system console log, with a low chance | 
| 1022 | that the user will ever see it. Moreover, as a genuine side-effect on global | |
| 1023 | process channels, there is no proper way to retract output when Isar command | |
| 40126 | 1024 | transactions are reset by the system. | 
| 39861 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 wenzelm parents: 
39859diff
changeset | 1025 |   \end{warn}
 | 
| 39872 | 1026 | |
| 1027 |   \begin{warn}
 | |
| 61854 | 1028 | The message channels should be used in a message-oriented manner. This means | 
| 1029 | that multi-line output that logically belongs together is issued by a single | |
| 69597 | 1030 | invocation of \<^ML>\<open>writeln\<close> etc.\ with the functional concatenation of all | 
| 61854 | 1031 | message constituents. | 
| 39872 | 1032 |   \end{warn}
 | 
| 58618 | 1033 | \<close> | 
| 1034 | ||
| 61854 | 1035 | text %mlex \<open> | 
| 1036 | The following example demonstrates a multi-line warning. Note that in some | |
| 1037 | situations the user sees only the first line, so the most important point | |
| 1038 | should be made first. | |
| 58618 | 1039 | \<close> | 
| 1040 | ||
| 1041 | ML_command \<open> | |
| 39872 | 1042 | warning (cat_lines | 
| 1043 | ["Beware the Jabberwock, my son!", | |
| 1044 | "The jaws that bite, the claws that catch!", | |
| 1045 | "Beware the Jubjub Bird, and shun", | |
| 1046 | "The frumious Bandersnatch!"]); | |
| 58618 | 1047 | \<close> | 
| 1048 | ||
| 59902 | 1049 | text \<open> | 
| 61416 | 1050 | \<^medskip> | 
| 61854 | 1051 | An alternative is to make a paragraph of freely-floating words as follows. | 
| 59902 | 1052 | \<close> | 
| 1053 | ||
| 1054 | ML_command \<open> | |
| 1055 | warning (Pretty.string_of (Pretty.para | |
| 1056 | "Beware the Jabberwock, my son! \ | |
| 1057 | \The jaws that bite, the claws that catch! \ | |
| 1058 | \Beware the Jubjub Bird, and shun \ | |
| 1059 | \The frumious Bandersnatch!")) | |
| 1060 | \<close> | |
| 1061 | ||
| 1062 | text \<open> | |
| 1063 | This has advantages with variable window / popup sizes, but might make it | |
| 1064 | harder to search for message content systematically, e.g.\ by other tools or | |
| 1065 | by humans expecting the ``verse'' of a formal message in a fixed layout. | |
| 1066 | \<close> | |
| 1067 | ||
| 58618 | 1068 | |
| 1069 | section \<open>Exceptions \label{sec:exceptions}\<close>
 | |
| 1070 | ||
| 61854 | 1071 | text \<open> | 
| 1072 | The Standard ML semantics of strict functional evaluation together with | |
| 78742 | 1073 | exceptions is rather well defined, but some fine points need to be observed | 
| 1074 | to avoid that ML programs go wrong despite static type-checking. | |
| 1075 | ||
| 1076 | Unlike official Standard ML, Isabelle/ML rejects catch-all patterns in | |
| 1077 | \<^ML_text>\<open>handle\<close> clauses: this improves the robustness of ML programs, | |
| 1078 | especially against arbitrary physical events (interrupts). | |
| 1079 | ||
| 61854 | 1080 | Exceptions in Isabelle/ML are subsequently categorized as follows. | 
| 1081 | \<close> | |
| 61506 | 1082 | |
| 1083 | paragraph \<open>Regular user errors.\<close> | |
| 61854 | 1084 | text \<open> | 
| 1085 | These are meant to provide informative feedback about malformed input etc. | |
| 1086 | ||
| 69597 | 1087 | The \<^emph>\<open>error\<close> function raises the corresponding \<^ML>\<open>ERROR\<close> exception, with a | 
| 1088 | plain text message as argument. \<^ML>\<open>ERROR\<close> exceptions can be handled | |
| 61854 | 1089 | internally, in order to be ignored, turned into other exceptions, or | 
| 1090 | cascaded by appending messages. If the corresponding Isabelle/Isar command | |
| 69597 | 1091 | terminates with an \<^ML>\<open>ERROR\<close> exception state, the system will print the | 
| 61854 | 1092 |   result on the error channel (see \secref{sec:message-channels}).
 | 
| 39854 | 1093 | |
| 61854 | 1094 | It is considered bad style to refer to internal function names or values in | 
| 1095 |   ML source notation in user error messages. Do not use \<open>@{make_string}\<close> nor
 | |
| 1096 |   \<open>@{here}\<close>!
 | |
| 39854 | 1097 | |
| 61854 | 1098 | Grammatical correctness of error messages can be improved by \<^emph>\<open>omitting\<close> | 
| 1099 | final punctuation: messages are often concatenated or put into a larger | |
| 1100 | context (e.g.\ augmented with source position). Note that punctuation after | |
| 1101 | formal entities (types, terms, theorems) is particularly prone to user | |
| 1102 | confusion. | |
| 61506 | 1103 | \<close> | 
| 1104 | ||
| 1105 | paragraph \<open>Program failures.\<close> | |
| 61854 | 1106 | text \<open> | 
| 1107 | There is a handful of standard exceptions that indicate general failure | |
| 78742 | 1108 | situations (e.g.\ \<^ML>\<open>Fail\<close>), or failures of core operations on logical | 
| 1109 |   entities (types, terms, theorems, theories, see \chref{ch:logic}).
 | |
| 39854 | 1110 | |
| 61854 | 1111 | These exceptions indicate a genuine breakdown of the program, so the main | 
| 78742 | 1112 | purpose is to determine quickly what has happened in the ML program. | 
| 1113 | Traditionally, the (short) exception message would include the name of an ML | |
| 1114 | function, although this is not strictly necessary, because the ML runtime | |
| 1115 | system attaches detailed source position stemming from the corresponding | |
| 1116 | \<^ML_text>\<open>raise\<close> keyword. | |
| 39854 | 1117 | |
| 61416 | 1118 | \<^medskip> | 
| 61854 | 1119 | User modules can always introduce their own custom exceptions locally, e.g.\ | 
| 1120 | to organize internal failures robustly without overlapping with existing | |
| 1121 | exceptions. Exceptions that are exposed in module signatures require extra | |
| 1122 | care, though, and should \<^emph>\<open>not\<close> be introduced by default. Surprise by users | |
| 1123 | of a module can be often minimized by using plain user errors instead. | |
| 61506 | 1124 | \<close> | 
| 1125 | ||
| 1126 | paragraph \<open>Interrupts.\<close> | |
| 61854 | 1127 | text \<open> | 
| 1128 | These indicate arbitrary system events: both the ML runtime system and the | |
| 78741 | 1129 | Isabelle/ML infrastructure may signal various exceptional situations by | 
| 1130 | raising special exceptions user code, satisfying the predicate | |
| 78681 | 1131 | \<^ML>\<open>Exn.is_interrupt\<close>. | 
| 57421 | 1132 | |
| 1133 | This is the one and only way that physical events can intrude an Isabelle/ML | |
| 1134 | program. Such an interrupt can mean out-of-memory, stack overflow, timeout, | |
| 1135 | internal signaling of threads, or a POSIX process signal. An Isabelle/ML | |
| 1136 | program that intercepts interrupts becomes dependent on physical effects of | |
| 78742 | 1137 | the environment (e.g.\ via \<^ML>\<open>Exn.capture\<close> without subsequent | 
| 1138 | \<^ML>\<open>Exn.release\<close>). | |
| 1139 | ||
| 1140 | Note that the original SML90 language had an \<^verbatim>\<open>Interrupt\<close> exception, but | |
| 1141 | that was excluded from SML97 to simplify ML the mathematical semantics. | |
| 1142 | Isabelle/ML does support physical interrupts thanks to special features of | |
| 1143 | the underlying Poly/ML compiler and runtime system. This works robustly, | |
| 1144 | because the old \<^ML_text>\<open>Interrupt\<close> constructor has been removed from the | |
| 1145 | ML environment, and catch-all patterns \<^ML_text>\<open>handle\<close> are rejected. | |
| 1146 | Thus user code becomes strictly transparent wrt.\ interrupts: physical | |
| 1147 | events are exposed to the toplevel, and the mathematical meaning of the | |
| 1148 | program becomes a partial function that is otherwise unchanged. | |
| 1149 | ||
| 1150 |   The Isabelle/ML antiquotation @{ML_antiquotation try}, with its syntactic
 | |
| 1151 | variants for \<^ML_text>\<open>catch\<close> or \<^ML_text>\<open>finally\<close>, supports | |
| 1152 | intermediate handling of interrupts and subsequent cleanup-operations, | |
| 1153 | without swallowing such physical event. | |
| 58618 | 1154 | \<close> | 
| 1155 | ||
| 1156 | text %mlref \<open> | |
| 39855 | 1157 |   \begin{mldecls}
 | 
| 73764 | 1158 |   @{define_ML try: "('a -> 'b) -> 'a -> 'b option"} \\
 | 
| 1159 |   @{define_ML can: "('a -> 'b) -> 'a -> bool"} \\
 | |
| 1160 |   @{define_ML_exception ERROR of string} \\
 | |
| 1161 |   @{define_ML_exception Fail of string} \\
 | |
| 1162 |   @{define_ML Exn.is_interrupt: "exn -> bool"} \\
 | |
| 1163 |   @{define_ML Exn.reraise: "exn -> 'a"} \\
 | |
| 1164 |   @{define_ML Runtime.exn_trace: "(unit -> 'a) -> 'a"} \\
 | |
| 39855 | 1165 |   \end{mldecls}
 | 
| 1166 | ||
| 78742 | 1167 | \<^descr> \<^ML>\<open>try\<close>~\<open>f x\<close> makes the partiality of evaluating \<open>f x\<close> explicit via | 
| 1168 | the option datatype. Interrupts are \<^emph>\<open>not\<close> handled here, i.e.\ this form | |
| 1169 | serves as safe replacement for the \<^emph>\<open>fragile\<close> version \<^ML_text>\<open>(SOME\<close>~\<open>f | |
| 69597 | 1170 | x\<close>~\<^ML_text>\<open>handle _ => NONE)\<close> that is occasionally seen in books about | 
| 78742 | 1171 | SML97. | 
| 39855 | 1172 | |
| 69597 | 1173 | \<^descr> \<^ML>\<open>can\<close> is similar to \<^ML>\<open>try\<close> with more abstract result. | 
| 1174 | ||
| 1175 | \<^descr> \<^ML>\<open>ERROR\<close>~\<open>msg\<close> represents user errors; this exception is normally | |
| 1176 | raised indirectly via the \<^ML>\<open>error\<close> function (see | |
| 61854 | 1177 |   \secref{sec:message-channels}).
 | 
| 39856 | 1178 | |
| 78742 | 1179 | \<^descr> \<^ML>\<open>Fail\<close>~\<open>msg\<close> represents general program failures, but not user errors. | 
| 1180 | ||
| 1181 | \<^descr> \<^ML>\<open>Exn.is_interrupt\<close> identifies interrupts, without mentioning | |
| 1182 | concrete exception constructors in user code. Since \<^ML_text>\<open>handle\<close> with | |
| 1183 | catch-all patterns is rejected, it cannot handle interrupts at all. In the | |
| 1184 | rare situations where this is really required, use \<^ML>\<open>Exn.capture\<close> and | |
| 1185 |   \<^ML>\<open>Exn.release\<close> (\secref{sec:managed-eval}).
 | |
| 61854 | 1186 | |
| 69597 | 1187 | \<^descr> \<^ML>\<open>Exn.reraise\<close>~\<open>exn\<close> raises exception \<open>exn\<close> while preserving its implicit | 
| 61854 | 1188 | position information (if possible, depending on the ML platform). | 
| 39856 | 1189 | |
| 69597 | 1190 | \<^descr> \<^ML>\<open>Runtime.exn_trace\<close>~\<^ML_text>\<open>(fn () =>\<close>~\<open>e\<close>\<^ML_text>\<open>)\<close> evaluates | 
| 61854 | 1191 | expression \<open>e\<close> while printing a full trace of its stack of nested exceptions | 
| 1192 | (if possible, depending on the ML platform). | |
| 39855 | 1193 | |
| 69597 | 1194 | Inserting \<^ML>\<open>Runtime.exn_trace\<close> into ML code temporarily is useful for | 
| 61854 | 1195 | debugging, but not suitable for production code. | 
| 58618 | 1196 | \<close> | 
| 1197 | ||
| 1198 | text %mlantiq \<open> | |
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39864diff
changeset | 1199 |   \begin{matharray}{rcl}
 | 
| 73552 | 1200 |   @{ML_antiquotation_def "try"} & : & \<open>ML_antiquotation\<close> \\
 | 
| 1201 |   @{ML_antiquotation_def "can"} & : & \<open>ML_antiquotation\<close> \\
 | |
| 61493 | 1202 |   @{ML_antiquotation_def "assert"} & : & \<open>ML_antiquotation\<close> \\
 | 
| 64465 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1203 |   @{ML_antiquotation_def "undefined"} & : & \<open>ML_antiquotation\<close> \\
 | 
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39864diff
changeset | 1204 |   \end{matharray}
 | 
| 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39864diff
changeset | 1205 | |
| 78741 | 1206 | \<^rail>\<open> | 
| 1207 |     (@@{ML_antiquotation try} | @@{ML_antiquotation can}) embedded
 | |
| 1208 | \<close> | |
| 1209 | ||
| 78742 | 1210 | |
| 1211 |   \<^descr> \<open>@{try}\<close> and \<open>{can}\<close> take embedded ML source as arguments, and modify the
 | |
| 1212 | evaluation analogously to the combinators \<^ML>\<open>try\<close> and \<^ML>\<open>can\<close> above, | |
| 1213 | but with special treatment of the interrupt state of the underlying ML | |
| 1214 | thread. There are also variants to support \<^verbatim>\<open>try_catch\<close> and \<^verbatim>\<open>try_finally\<close> | |
| 1215 | blocks similar to Scala. | |
| 1216 | ||
| 1217 | The substructure of the embedded argument supports the following syntax | |
| 1218 | variants: | |
| 1219 | ||
| 1220 | \<^rail>\<open> | |
| 1221 |     @{syntax_def try_catch}: @{syntax expr} @'catch' @{syntax handler};
 | |
| 1222 |     @{syntax_def try_finally}: @{syntax expr} @'finally' @{syntax cleanup};
 | |
| 1223 |     @{syntax_def try}: @{syntax expr};
 | |
| 1224 |     @{syntax_def can}: @{syntax expr}
 | |
| 1225 | \<close> | |
| 1226 | ||
| 1227 |   The @{syntax handler} of \<^verbatim>\<open>try_catch\<close> follows the syntax of \<^ML_text>\<open>fn\<close>
 | |
| 1228 | patterns, so it is similar to \<^ML_text>\<open>handle\<close>: the key difference is | |
| 1229 | that interrupts are always passed-through via \<^ML>\<open>Exn.reraise\<close>. | |
| 1230 | ||
| 1231 |   The @{syntax cleanup} expression of \<^verbatim>\<open>try_finally\<close> is always invoked,
 | |
| 1232 | regardless of the overall exception result, and afterwards exceptions are | |
| 1233 | passed-through via \<^ML>\<open>Exn.reraise\<close>. | |
| 1234 | ||
| 1235 |   Both the @{syntax handler} and @{syntax cleanup} are evaluated with further
 | |
| 1236 | interrupts disabled! These expressions should terminate promptly; timeouts | |
| 1237 | don't work here. | |
| 1238 | ||
| 1239 | \<^medskip> | |
| 1240 | Implementation details can be seen in \<^ML>\<open>Isabelle_Thread.try_catch\<close>, | |
| 1241 | \<^ML>\<open>Isabelle_Thread.try_finally\<close>, \<^ML>\<open>Isabelle_Thread.try\<close>, and | |
| 1242 | \<^ML>\<open>Isabelle_Thread.can\<close>, respectively. The ML antiquotations add | |
| 1243 | functional abstractions as required for these ``special forms'' of | |
| 1244 | Isabelle/ML. | |
| 73552 | 1245 | |
| 1246 |   \<^descr> \<open>@{assert}\<close> inlines a function \<^ML_type>\<open>bool -> unit\<close> that raises
 | |
| 1247 | \<^ML>\<open>Fail\<close> if the argument is \<^ML>\<open>false\<close>. Due to inlining the source | |
| 1248 | position of failed assertions is included in the error output. | |
| 64465 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1249 | |
| 69597 | 1250 |   \<^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: 
64464diff
changeset | 1251 | behaves as in some function application of an undefined case. | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1252 | \<close> | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1253 | |
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1254 | text %mlex \<open> | 
| 78742 | 1255 | |
| 1256 | We define total versions of division: any failures are swept under the | |
| 73552 | 1257 | carpet and mapped to a default value. Thus division-by-zero becomes 0, but | 
| 78742 | 1258 | there could be other exceptions like overflow that produce the same result. | 
| 1259 | ||
| 1260 | For unbounded integers such side-errors do not happen, but it might still be | |
| 1261 | better to be explicit about exception patterns (second version below). | |
| 73552 | 1262 | \<close> | 
| 1263 | ||
| 1264 | ML \<open> | |
| 78742 | 1265 | fun div_total1 x y = \<^try>\<open>x div y catch _ => 0\<close>; | 
| 1266 | fun div_total2 x y = \<^try>\<open>x div y catch Div => 0\<close>; | |
| 1267 | ||
| 1268 | \<^assert> (div_total1 1 0 = 0); | |
| 1269 | \<^assert> (div_total2 1 0 = 0); | |
| 73552 | 1270 | \<close> | 
| 1271 | ||
| 1272 | text \<open> | |
| 69597 | 1273 | 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: 
64464diff
changeset | 1274 | as follows: | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1275 | \<close> | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1276 | |
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1277 | ML \<open>fun undefined _ = raise Match\<close> | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1278 | |
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1279 | text \<open> | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1280 | \<^medskip> | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1281 |   The following variant uses the antiquotation @{ML_antiquotation undefined}
 | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1282 | instead: | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1283 | \<close> | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1284 | |
| 78741 | 1285 | ML \<open>fun undefined _ = @{undefined}\<close>
 | 
| 64465 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1286 | |
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1287 | text \<open> | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1288 | \<^medskip> | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1289 | Here is the same, using control-symbol notation for the antiquotation, with | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1290 | special rendering of \<^verbatim>\<open>\<^undefined>\<close>: | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1291 | \<close> | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1292 | |
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1293 | ML \<open>fun undefined _ = \<^undefined>\<close> | 
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1294 | |
| 
73069f272f42
documentation of @{undefined} (actually introduced in Isabelle2016);
 wenzelm parents: 
64464diff
changeset | 1295 | text \<open> | 
| 78741 | 1296 | \<^medskip> | 
| 1297 | Semantically, all forms are equivalent to a function definition without any | |
| 1298 | clauses, but that is syntactically not allowed in ML. | |
| 58618 | 1299 | \<close> | 
| 1300 | ||
| 1301 | ||
| 1302 | section \<open>Strings of symbols \label{sec:symbols}\<close>
 | |
| 1303 | ||
| 61854 | 1304 | text \<open> | 
| 1305 | A \<^emph>\<open>symbol\<close> constitutes the smallest textual unit in Isabelle/ML --- raw ML | |
| 1306 | characters are normally not encountered at all. Isabelle strings consist of | |
| 1307 | a sequence of symbols, represented as a packed string or an exploded list of | |
| 1308 | strings. Each symbol is in itself a small string, which has either one of | |
| 61504 | 1309 | the following forms: | 
| 1310 | ||
| 1311 | \<^enum> a single ASCII character ``\<open>c\<close>'', for example ``\<^verbatim>\<open>a\<close>'', | |
| 1312 | ||
| 1313 | \<^enum> a codepoint according to UTF-8 (non-ASCII byte sequence), | |
| 1314 | ||
| 1315 | \<^enum> a regular symbol ``\<^verbatim>\<open>\<ident>\<close>'', for example ``\<^verbatim>\<open>\<alpha>\<close>'', | |
| 1316 | ||
| 1317 | \<^enum> a control symbol ``\<^verbatim>\<open>\<^ident>\<close>'', for example ``\<^verbatim>\<open>\<^bold>\<close>'', | |
| 1318 | ||
| 1319 | The \<open>ident\<close> syntax for symbol names is \<open>letter (letter | digit)\<^sup>*\<close>, where | |
| 1320 | \<open>letter = A..Za..z\<close> and \<open>digit = 0..9\<close>. There are infinitely many regular | |
| 1321 | symbols and control symbols, but a fixed collection of standard symbols is | |
| 1322 | treated specifically. For example, ``\<^verbatim>\<open>\<alpha>\<close>'' is classified as a letter, which | |
| 1323 | means it may occur within regular Isabelle identifiers. | |
| 52421 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1324 | |
| 57421 | 1325 | The character set underlying Isabelle symbols is 7-bit ASCII, but 8-bit | 
| 1326 | character sequences are passed-through unchanged. Unicode/UCS data in UTF-8 | |
| 1327 | encoding is processed in a non-strict fashion, such that well-formed code | |
| 1328 | sequences are recognized accordingly. Unicode provides its own collection of | |
| 1329 | mathematical symbols, but within the core Isabelle/ML world there is no link | |
| 1330 | to the standard collection of Isabelle regular symbols. | |
| 1331 | ||
| 61416 | 1332 | \<^medskip> | 
| 61504 | 1333 | Output of Isabelle symbols depends on the print mode. For example, the | 
| 1334 |   standard {\LaTeX} setup of the Isabelle document preparation system would
 | |
| 1335 | present ``\<^verbatim>\<open>\<alpha>\<close>'' as \<open>\<alpha>\<close>, and ``\<^verbatim>\<open>\<^bold>\<alpha>\<close>'' as \<open>\<^bold>\<alpha>\<close>. On-screen rendering usually | |
| 1336 | works by mapping a finite subset of Isabelle symbols to suitable Unicode | |
| 1337 | characters. | |
| 58618 | 1338 | \<close> | 
| 1339 | ||
| 1340 | text %mlref \<open> | |
| 52421 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1341 |   \begin{mldecls}
 | 
| 73764 | 1342 |   @{define_ML_type Symbol.symbol = string} \\
 | 
| 1343 |   @{define_ML Symbol.explode: "string -> Symbol.symbol list"} \\
 | |
| 1344 |   @{define_ML Symbol.is_letter: "Symbol.symbol -> bool"} \\
 | |
| 1345 |   @{define_ML Symbol.is_digit: "Symbol.symbol -> bool"} \\
 | |
| 1346 |   @{define_ML Symbol.is_quasi: "Symbol.symbol -> bool"} \\
 | |
| 1347 |   @{define_ML Symbol.is_blank: "Symbol.symbol -> bool"} \\
 | |
| 52421 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1348 |   \end{mldecls}
 | 
| 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1349 |   \begin{mldecls}
 | 
| 73764 | 1350 |   @{define_ML_type "Symbol.sym"} \\
 | 
| 1351 |   @{define_ML Symbol.decode: "Symbol.symbol -> Symbol.sym"} \\
 | |
| 52421 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1352 |   \end{mldecls}
 | 
| 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1353 | |
| 69597 | 1354 | \<^descr> Type \<^ML_type>\<open>Symbol.symbol\<close> represents individual Isabelle symbols. | 
| 1355 | ||
| 1356 | \<^descr> \<^ML>\<open>Symbol.explode\<close>~\<open>str\<close> produces a symbol list from the packed form. | |
| 1357 | This function supersedes \<^ML>\<open>String.explode\<close> for virtually all purposes | |
| 61854 | 1358 | of manipulating text in Isabelle!\<^footnote>\<open>The runtime overhead for exploded strings | 
| 1359 | is mainly that of the list structure: individual symbols that happen to be a | |
| 1360 | singleton string do not require extra memory in Poly/ML.\<close> | |
| 52421 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1361 | |
| 75617 | 1362 | \<^descr> \<^ML>\<open>Symbol.is_letter\<close>, \<^ML>\<open>Symbol.is_digit\<close>, | 
| 1363 | \<^ML>\<open>Symbol.is_quasi\<close>, \<^ML>\<open>Symbol.is_blank\<close> classify standard symbols | |
| 76987 | 1364 | according to fixed syntactic conventions of Isabelle, cf.\ \<^cite>\<open>"isabelle-isar-ref"\<close>. | 
| 52421 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1365 | |
| 69597 | 1366 | \<^descr> Type \<^ML_type>\<open>Symbol.sym\<close> is a concrete datatype that represents the | 
| 75617 | 1367 | different kinds of symbols explicitly, with constructors | 
| 1368 | \<^ML>\<open>Symbol.Char\<close>, \<^ML>\<open>Symbol.UTF8\<close>, \<^ML>\<open>Symbol.Sym\<close>, | |
| 1369 | \<^ML>\<open>Symbol.Control\<close>, \<^ML>\<open>Symbol.Malformed\<close>. | |
| 69597 | 1370 | |
| 1371 | \<^descr> \<^ML>\<open>Symbol.decode\<close> converts the string representation of a symbol into | |
| 61854 | 1372 | the datatype version. | 
| 61506 | 1373 | \<close> | 
| 1374 | ||
| 1375 | paragraph \<open>Historical note.\<close> | |
| 61854 | 1376 | text \<open> | 
| 69597 | 1377 | In the original SML90 standard the primitive ML type \<^ML_type>\<open>char\<close> did not | 
| 1378 | exists, and \<^ML_text>\<open>explode: string -> string list\<close> produced a list of | |
| 1379 | singleton strings like \<^ML>\<open>raw_explode: string -> string list\<close> in | |
| 61506 | 1380 | Isabelle/ML today. When SML97 came out, Isabelle did not adopt its somewhat | 
| 1381 | anachronistic 8-bit or 16-bit characters, but the idea of exploding a string | |
| 1382 | into a list of small strings was extended to ``symbols'' as explained above. | |
| 1383 | Thus Isabelle sources can refer to an infinite store of user-defined | |
| 1384 | symbols, without having to worry about the multitude of Unicode encodings | |
| 61854 | 1385 | that have emerged over the years. | 
| 1386 | \<close> | |
| 58618 | 1387 | |
| 1388 | ||
| 1389 | section \<open>Basic data types\<close> | |
| 1390 | ||
| 61854 | 1391 | text \<open> | 
| 1392 | The basis library proposal of SML97 needs to be treated with caution. Many | |
| 1393 | of its operations simply do not fit with important Isabelle/ML conventions | |
| 1394 | (like ``canonical argument order'', see | |
| 1395 |   \secref{sec:canonical-argument-order}), others cause problems with the
 | |
| 75617 | 1396 | parallel evaluation model of Isabelle/ML (such as \<^ML>\<open>TextIO.print\<close> or | 
| 1397 | \<^ML>\<open>OS.Process.system\<close>). | |
| 39859 | 1398 | |
| 61854 | 1399 | Subsequently we give a brief overview of important operations on basic ML | 
| 1400 | data types. | |
| 58618 | 1401 | \<close> | 
| 1402 | ||
| 1403 | ||
| 1404 | subsection \<open>Characters\<close> | |
| 1405 | ||
| 1406 | text %mlref \<open> | |
| 39863 | 1407 |   \begin{mldecls}
 | 
| 73764 | 1408 |   @{define_ML_type char} \\
 | 
| 39863 | 1409 |   \end{mldecls}
 | 
| 1410 | ||
| 69597 | 1411 | \<^descr> Type \<^ML_type>\<open>char\<close> is \<^emph>\<open>not\<close> used. The smallest textual unit in Isabelle | 
| 61854 | 1412 |   is represented as a ``symbol'' (see \secref{sec:symbols}).
 | 
| 58618 | 1413 | \<close> | 
| 1414 | ||
| 1415 | ||
| 1416 | subsection \<open>Strings\<close> | |
| 1417 | ||
| 1418 | text %mlref \<open> | |
| 52421 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1419 |   \begin{mldecls}
 | 
| 73764 | 1420 |   @{define_ML_type string} \\
 | 
| 52421 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1421 |   \end{mldecls}
 | 
| 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1422 | |
| 69597 | 1423 | \<^descr> Type \<^ML_type>\<open>string\<close> represents immutable vectors of 8-bit characters. | 
| 61854 | 1424 | There are operations in SML to convert back and forth to actual byte | 
| 1425 | vectors, which are seldom used. | |
| 52421 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1426 | |
| 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1427 | This historically important raw text representation is used for | 
| 61854 | 1428 | Isabelle-specific purposes with the following implicit substructures packed | 
| 1429 | into the string content: | |
| 52421 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1430 | |
| 75617 | 1431 |     \<^enum> sequence of Isabelle symbols (see also \secref{sec:symbols}), with
 | 
| 1432 | \<^ML>\<open>Symbol.explode\<close> as key operation; | |
| 61458 | 1433 | |
| 76987 | 1434 | \<^enum> XML tree structure via YXML (see also \<^cite>\<open>"isabelle-system"\<close>), with | 
| 69597 | 1435 | \<^ML>\<open>YXML.parse_body\<close> as key operation. | 
| 52421 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1436 | |
| 58723 | 1437 | Note that Isabelle/ML string literals may refer Isabelle symbols like | 
| 61854 | 1438 | ``\<^verbatim>\<open>\<alpha>\<close>'' natively, \<^emph>\<open>without\<close> escaping the backslash. This is a consequence | 
| 1439 | of Isabelle treating all source text as strings of symbols, instead of raw | |
| 1440 | characters. | |
| 75619 | 1441 | |
| 1442 | ||
| 1443 |   \begin{warn}
 | |
| 1444 | The regular \<^verbatim>\<open>64_32\<close> platform of Poly/ML has a size limit of 64\,MB for | |
| 1445 | \<^ML_type>\<open>string\<close> values. This is usually sufficient for text | |
| 1446 | applications, with a little bit of YXML markup. Very large XML trees or | |
| 1447 | binary blobs are better stored as scalable byte strings, see type | |
| 1448 | \<^ML_type>\<open>Bytes.T\<close> and corresponding operations in | |
| 1449 | \<^file>\<open>~~/src/Pure/General/bytes.ML\<close>. | |
| 1450 |   \end{warn}
 | |
| 58618 | 1451 | \<close> | 
| 1452 | ||
| 61854 | 1453 | text %mlex \<open> | 
| 1454 | The subsequent example illustrates the difference of physical addressing of | |
| 1455 | bytes versus logical addressing of symbols in Isabelle strings. | |
| 58618 | 1456 | \<close> | 
| 1457 | ||
| 1458 | ML_val \<open> | |
| 52421 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1459 | val s = "\<A>"; | 
| 
6d93140a206c
clarified strings of symbols, including ML string literals;
 wenzelm parents: 
52420diff
changeset | 1460 | |
| 69597 | 1461 | \<^assert> (length (Symbol.explode s) = 1); | 
| 1462 | \<^assert> (size s = 4); | |
| 58618 | 1463 | \<close> | 
| 1464 | ||
| 61854 | 1465 | text \<open> | 
| 1466 | Note that in Unicode renderings of the symbol \<open>\<A>\<close>, variations of encodings | |
| 1467 | like UTF-8 or UTF-16 pose delicate questions about the multi-byte | |
| 1468 | representations of its codepoint, which is outside of the 16-bit address | |
| 1469 | space of the original Unicode standard from the 1990-ies. In Isabelle/ML it | |
| 1470 | is just ``\<^verbatim>\<open>\<A>\<close>'' literally, using plain ASCII characters beyond any | |
| 1471 | doubts. | |
| 1472 | \<close> | |
| 58618 | 1473 | |
| 1474 | ||
| 1475 | subsection \<open>Integers\<close> | |
| 1476 | ||
| 1477 | text %mlref \<open> | |
| 39862 | 1478 |   \begin{mldecls}
 | 
| 73764 | 1479 |   @{define_ML_type int} \\
 | 
| 39862 | 1480 |   \end{mldecls}
 | 
| 1481 | ||
| 69597 | 1482 | \<^descr> Type \<^ML_type>\<open>int\<close> represents regular mathematical integers, which are | 
| 61854 | 1483 | \<^emph>\<open>unbounded\<close>. Overflow is treated properly, but should never happen in | 
| 1484 | practice.\<^footnote>\<open>The size limit for integer bit patterns in memory is 64\,MB for | |
| 75619 | 1485 | the regular \<^verbatim>\<open>64_32\<close> platform, and much higher for native \<^verbatim>\<open>64\<close> | 
| 1486 | architecture.\<close> | |
| 39862 | 1487 | |
| 69597 | 1488 | Structure \<^ML_structure>\<open>IntInf\<close> of SML97 is obsolete and superseded by | 
| 1489 | \<^ML_structure>\<open>Int\<close>. Structure \<^ML_structure>\<open>Integer\<close> in | |
| 63680 | 1490 | \<^file>\<open>~~/src/Pure/General/integer.ML\<close> provides some additional operations. | 
| 58618 | 1491 | \<close> | 
| 1492 | ||
| 1493 | ||
| 63215 | 1494 | subsection \<open>Rational numbers\<close> | 
| 1495 | ||
| 1496 | text %mlref \<open> | |
| 1497 |   \begin{mldecls}
 | |
| 73764 | 1498 |   @{define_ML_type Rat.rat} \\
 | 
| 63215 | 1499 |   \end{mldecls}
 | 
| 1500 | ||
| 69597 | 1501 | \<^descr> Type \<^ML_type>\<open>Rat.rat\<close> represents rational numbers, based on the | 
| 63215 | 1502 | unbounded integers of Poly/ML. | 
| 1503 | ||
| 1504 | Literal rationals may be written with special antiquotation syntax | |
| 1505 | \<^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 | |
| 1506 | \<^verbatim>\<open>@~1/4\<close> or \<^verbatim>\<open>@10\<close>. The ML toplevel pretty printer uses the same format. | |
| 1507 | ||
| 1508 | Standard operations are provided via ad-hoc overloading of \<^verbatim>\<open>+\<close>, \<^verbatim>\<open>-\<close>, \<^verbatim>\<open>*\<close>, | |
| 1509 | \<^verbatim>\<open>/\<close>, etc. | |
| 1510 | \<close> | |
| 1511 | ||
| 1512 | ||
| 58618 | 1513 | subsection \<open>Time\<close> | 
| 1514 | ||
| 1515 | text %mlref \<open> | |
| 40302 | 1516 |   \begin{mldecls}
 | 
| 73764 | 1517 |   @{define_ML_type Time.time} \\
 | 
| 1518 |   @{define_ML seconds: "real -> Time.time"} \\
 | |
| 40302 | 1519 |   \end{mldecls}
 | 
| 1520 | ||
| 69597 | 1521 | \<^descr> Type \<^ML_type>\<open>Time.time\<close> represents time abstractly according to the | 
| 61854 | 1522 | SML97 basis library definition. This is adequate for internal ML operations, | 
| 1523 | but awkward in concrete time specifications. | |
| 40302 | 1524 | |
| 69597 | 1525 | \<^descr> \<^ML>\<open>seconds\<close>~\<open>s\<close> turns the concrete scalar \<open>s\<close> (measured in seconds) into | 
| 61854 | 1526 | an abstract time value. Floating point numbers are easy to use as | 
| 1527 |   configuration options in the context (see \secref{sec:config-options}) or
 | |
| 1528 | system options that are maintained externally. | |
| 58618 | 1529 | \<close> | 
| 1530 | ||
| 1531 | ||
| 1532 | subsection \<open>Options\<close> | |
| 1533 | ||
| 1534 | text %mlref \<open> | |
| 39859 | 1535 |   \begin{mldecls}
 | 
| 73764 | 1536 |   @{define_ML Option.map: "('a -> 'b) -> 'a option -> 'b option"} \\
 | 
| 1537 |   @{define_ML is_some: "'a option -> bool"} \\
 | |
| 1538 |   @{define_ML is_none: "'a option -> bool"} \\
 | |
| 1539 |   @{define_ML the: "'a option -> 'a"} \\
 | |
| 1540 |   @{define_ML these: "'a list option -> 'a list"} \\
 | |
| 1541 |   @{define_ML the_list: "'a option -> 'a list"} \\
 | |
| 1542 |   @{define_ML the_default: "'a -> 'a option -> 'a"} \\
 | |
| 39859 | 1543 |   \end{mldecls}
 | 
| 77907 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1544 | |
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1545 |   \begin{matharray}{rcl}
 | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1546 |   @{ML_antiquotation_def "if_none"} & : & \<open>ML_antiquotation\<close> \\
 | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1547 |   \end{matharray}
 | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1548 | |
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1549 |   \<^rail>\<open>@@{ML_antiquotation if_none} embedded\<close>
 | 
| 58618 | 1550 | \<close> | 
| 1551 | ||
| 61854 | 1552 | text \<open> | 
| 69597 | 1553 | Apart from \<^ML>\<open>Option.map\<close> most other operations defined in structure | 
| 1554 | \<^ML_structure>\<open>Option\<close> are alien to Isabelle/ML and never used. The | |
| 63680 | 1555 | operations shown above are defined in \<^file>\<open>~~/src/Pure/General/basics.ML\<close>. | 
| 77907 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1556 | |
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1557 | Note that the function \<^ML>\<open>the_default\<close> is strict in all of its | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1558 | arguments, the default value is evaluated beforehand, even if not required | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1559 |   later. In contrast, the antiquotation @{ML_antiquotation "if_none"} is
 | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1560 | non-strict: the given expression is only evaluated for an application to | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1561 | \<^ML>\<open>NONE\<close>. This allows to work with exceptions like this: | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1562 | \<close> | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1563 | |
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1564 | ML \<open> | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1565 | fun div_total x y = | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1566 | \<^try>\<open>x div y\<close> |> the_default 0; | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1567 | |
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1568 | fun div_error x y = | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1569 | \<^try>\<open>x div y\<close> |> \<^if_none>\<open>error "Division by zero"\<close>; | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1570 | \<close> | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1571 | |
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1572 | text \<open> | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1573 | Of course, it is also possible to handle exceptions directly, without an | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1574 | intermediate option construction: | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1575 | \<close> | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1576 | |
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1577 | ML \<open> | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1578 | fun div_total x y = | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1579 | x div y handle Div => 0; | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1580 | |
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1581 | fun div_error x y = | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1582 | x div y handle Div => error "Division by zero"; | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1583 | \<close> | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1584 | |
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1585 | text \<open> | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1586 | The first form works better in longer chains of functional composition, with | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1587 | combinators like \<^ML>\<open>|>\<close> or \<^ML>\<open>#>\<close> or \<^ML>\<open>o\<close>. The second form is more | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1588 | adequate in elementary expressions: there is no need to pretend that | 
| 
ee9785abbcd6
provide ML antiquotation "if_none": non-strict version of "the_default";
 wenzelm parents: 
76987diff
changeset | 1589 | Isabelle/ML is actually a version of Haskell. | 
| 61854 | 1590 | \<close> | 
| 58618 | 1591 | |
| 1592 | ||
| 1593 | subsection \<open>Lists\<close> | |
| 1594 | ||
| 61854 | 1595 | text \<open> | 
| 1596 | Lists are ubiquitous in ML as simple and light-weight ``collections'' for | |
| 1597 | many everyday programming tasks. Isabelle/ML provides important additions | |
| 1598 | and improvements over operations that are predefined in the SML97 library. | |
| 1599 | \<close> | |
| 58618 | 1600 | |
| 1601 | text %mlref \<open> | |
| 39863 | 1602 |   \begin{mldecls}
 | 
| 73764 | 1603 |   @{define_ML cons: "'a -> 'a list -> 'a list"} \\
 | 
| 1604 |   @{define_ML member: "('b * 'a -> bool) -> 'a list -> 'b -> bool"} \\
 | |
| 1605 |   @{define_ML insert: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
 | |
| 1606 |   @{define_ML remove: "('b * 'a -> bool) -> 'b -> 'a list -> 'a list"} \\
 | |
| 1607 |   @{define_ML update: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
 | |
| 39863 | 1608 |   \end{mldecls}
 | 
| 1609 | ||
| 69597 | 1610 | \<^descr> \<^ML>\<open>cons\<close>~\<open>x xs\<close> evaluates to \<open>x :: xs\<close>. | 
| 39863 | 1611 | |
| 61854 | 1612 | Tupled infix operators are a historical accident in Standard ML. The curried | 
| 69597 | 1613 | \<^ML>\<open>cons\<close> amends this, but it should be only used when partial application | 
| 61854 | 1614 | is required. | 
| 39863 | 1615 | |
| 69597 | 1616 | \<^descr> \<^ML>\<open>member\<close>, \<^ML>\<open>insert\<close>, \<^ML>\<open>remove\<close>, \<^ML>\<open>update\<close> treat lists as a | 
| 63680 | 1617 | set-like container that maintains the order of elements. See | 
| 1618 | \<^file>\<open>~~/src/Pure/library.ML\<close> for the full specifications (written in ML). | |
| 69597 | 1619 | There are some further derived operations like \<^ML>\<open>union\<close> or \<^ML>\<open>inter\<close>. | 
| 1620 | ||
| 1621 | Note that \<^ML>\<open>insert\<close> is conservative about elements that are already a | |
| 1622 | \<^ML>\<open>member\<close> of the list, while \<^ML>\<open>update\<close> ensures that the latest entry | |
| 61854 | 1623 | is always put in front. The latter discipline is often more appropriate in | 
| 1624 |   declarations of context data (\secref{sec:context-data}) that are issued by
 | |
| 1625 | the user in Isar source: later declarations take precedence over earlier | |
| 1626 | ones. \<close> | |
| 58618 | 1627 | |
| 61854 | 1628 | text %mlex \<open> | 
| 69597 | 1629 | Using canonical \<^ML>\<open>fold\<close> together with \<^ML>\<open>cons\<close> (or similar standard | 
| 61854 | 1630 | operations) alternates the orientation of data. The is quite natural and | 
| 69597 | 1631 | should not be altered forcible by inserting extra applications of \<^ML>\<open>rev\<close>. | 
| 1632 | The alternative \<^ML>\<open>fold_rev\<close> can be used in the few situations, where | |
| 61854 | 1633 | alternation should be prevented. | 
| 58618 | 1634 | \<close> | 
| 1635 | ||
| 59902 | 1636 | ML_val \<open> | 
| 39863 | 1637 | val items = [1, 2, 3, 4, 5, 6, 7, 8, 9, 10]; | 
| 1638 | ||
| 1639 | val list1 = fold cons items []; | |
| 69597 | 1640 | \<^assert> (list1 = rev items); | 
| 39866 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 wenzelm parents: 
39864diff
changeset | 1641 | |
| 39863 | 1642 | val list2 = fold_rev cons items []; | 
| 69597 | 1643 | \<^assert> (list2 = items); | 
| 58618 | 1644 | \<close> | 
| 1645 | ||
| 61854 | 1646 | text \<open> | 
| 1647 | The subsequent example demonstrates how to \<^emph>\<open>merge\<close> two lists in a natural | |
| 1648 | way. | |
| 1649 | \<close> | |
| 58618 | 1650 | |
| 59902 | 1651 | ML_val \<open> | 
| 39883 | 1652 | fun merge_lists eq (xs, ys) = fold_rev (insert eq) ys xs; | 
| 58618 | 1653 | \<close> | 
| 1654 | ||
| 61854 | 1655 | text \<open> | 
| 1656 | Here the first list is treated conservatively: only the new elements from | |
| 75617 | 1657 | the second list are inserted. The inside-out order of insertion via | 
| 1658 | \<^ML>\<open>fold_rev\<close> attempts to preserve the order of elements in the result. | |
| 39883 | 1659 | |
| 1660 | This way of merging lists is typical for context data | |
| 69597 | 1661 |   (\secref{sec:context-data}). See also \<^ML>\<open>merge\<close> as defined in
 | 
| 63680 | 1662 | \<^file>\<open>~~/src/Pure/library.ML\<close>. | 
| 58618 | 1663 | \<close> | 
| 1664 | ||
| 1665 | ||
| 1666 | subsection \<open>Association lists\<close> | |
| 1667 | ||
| 61854 | 1668 | text \<open> | 
| 1669 | The operations for association lists interpret a concrete list of pairs as a | |
| 1670 | finite function from keys to values. Redundant representations with multiple | |
| 1671 | occurrences of the same key are implicitly normalized: lookup and update | |
| 1672 | only take the first occurrence into account. | |
| 58618 | 1673 | \<close> | 
| 1674 | ||
| 1675 | text \<open> | |
| 39875 
648c930125f6
more on "Association lists", based on more succinct version of older material;
 wenzelm parents: 
39874diff
changeset | 1676 |   \begin{mldecls}
 | 
| 73764 | 1677 |   @{define_ML AList.lookup: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option"} \\
 | 
| 1678 |   @{define_ML AList.defined: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool"} \\
 | |
| 1679 |   @{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: 
39874diff
changeset | 1680 |   \end{mldecls}
 | 
| 
648c930125f6
more on "Association lists", based on more succinct version of older material;
 wenzelm parents: 
39874diff
changeset | 1681 | |
| 69597 | 1682 | \<^descr> \<^ML>\<open>AList.lookup\<close>, \<^ML>\<open>AList.defined\<close>, \<^ML>\<open>AList.update\<close> implement the | 
| 61854 | 1683 | main ``framework operations'' for mappings in Isabelle/ML, following | 
| 1684 | standard conventions for their names and types. | |
| 39875 
648c930125f6
more on "Association lists", based on more succinct version of older material;
 wenzelm parents: 
39874diff
changeset | 1685 | |
| 61854 | 1686 | Note that a function called \<^verbatim>\<open>lookup\<close> is obliged to express its partiality | 
| 1687 | via an explicit option element. There is no choice to raise an exception, | |
| 1688 | without changing the name to something like \<open>the_element\<close> or \<open>get\<close>. | |
| 61493 | 1689 | |
| 69597 | 1690 | The \<open>defined\<close> operation is essentially a contraction of \<^ML>\<open>is_some\<close> and | 
| 61854 | 1691 | \<^verbatim>\<open>lookup\<close>, but this is sufficiently frequent to justify its independent | 
| 1692 | existence. This also gives the implementation some opportunity for peep-hole | |
| 1693 | optimization. | |
| 39875 
648c930125f6
more on "Association lists", based on more succinct version of older material;
 wenzelm parents: 
39874diff
changeset | 1694 | |
| 
648c930125f6
more on "Association lists", based on more succinct version of older material;
 wenzelm parents: 
39874diff
changeset | 1695 | |
| 57421 | 1696 | Association lists are adequate as simple implementation of finite mappings | 
| 1697 | in many practical situations. A more advanced table structure is defined in | |
| 63680 | 1698 | \<^file>\<open>~~/src/Pure/General/table.ML\<close>; that version scales easily to thousands or | 
| 1699 | millions of elements. | |
| 58618 | 1700 | \<close> | 
| 1701 | ||
| 1702 | ||
| 1703 | subsection \<open>Unsynchronized references\<close> | |
| 1704 | ||
| 1705 | text %mlref \<open> | |
| 39859 | 1706 |   \begin{mldecls}
 | 
| 73769 | 1707 |   @{define_ML_type 'a "Unsynchronized.ref"} \\
 | 
| 73764 | 1708 |   @{define_ML Unsynchronized.ref: "'a -> 'a Unsynchronized.ref"} \\
 | 
| 1709 |   @{define_ML "!": "'a Unsynchronized.ref -> 'a"} \\
 | |
| 1710 |   @{define_ML_infix ":=" : "'a Unsynchronized.ref * 'a -> unit"} \\
 | |
| 39859 | 1711 |   \end{mldecls}
 | 
| 58618 | 1712 | \<close> | 
| 1713 | ||
| 61854 | 1714 | text \<open> | 
| 1715 | Due to ubiquitous parallelism in Isabelle/ML (see also | |
| 1716 |   \secref{sec:multi-threading}), the mutable reference cells of Standard ML
 | |
| 1717 | are notorious for causing problems. In a highly parallel system, both | |
| 1718 | correctness \<^emph>\<open>and\<close> performance are easily degraded when using mutable data. | |
| 39859 | 1719 | |
| 69597 | 1720 | The unwieldy name of \<^ML>\<open>Unsynchronized.ref\<close> for the constructor for | 
| 61854 | 1721 | references in Isabelle/ML emphasizes the inconveniences caused by | 
| 73765 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 wenzelm parents: 
73764diff
changeset | 1722 | mutability. Existing operations \<^ML>\<open>!\<close> and \<^ML_infix>\<open>:=\<close> are unchanged, | 
| 61854 | 1723 | but should be used with special precautions, say in a strictly local | 
| 1724 | situation that is guaranteed to be restricted to sequential evaluation --- | |
| 1725 | now and in the future. | |
| 40508 | 1726 | |
| 1727 |   \begin{warn}
 | |
| 69597 | 1728 | Never \<^ML_text>\<open>open Unsynchronized\<close>, not even in a local scope! | 
| 40508 | 1729 | Pretending that mutable state is no problem is a very bad idea. | 
| 1730 |   \end{warn}
 | |
| 58618 | 1731 | \<close> | 
| 1732 | ||
| 1733 | ||
| 1734 | section \<open>Thread-safe programming \label{sec:multi-threading}\<close>
 | |
| 1735 | ||
| 61854 | 1736 | text \<open> | 
| 1737 | Multi-threaded execution has become an everyday reality in Isabelle since | |
| 1738 | Poly/ML 5.2.1 and Isabelle2008. Isabelle/ML provides implicit and explicit | |
| 1739 | parallelism by default, and there is no way for user-space tools to ``opt | |
| 1740 | out''. ML programs that are purely functional, output messages only via the | |
| 1741 |   official channels (\secref{sec:message-channels}), and do not intercept
 | |
| 1742 |   interrupts (\secref{sec:exceptions}) can participate in the multi-threaded
 | |
| 39868 | 1743 | environment immediately without further ado. | 
| 1744 | ||
| 61854 | 1745 | More ambitious tools with more fine-grained interaction with the environment | 
| 1746 | need to observe the principles explained below. | |
| 58618 | 1747 | \<close> | 
| 1748 | ||
| 1749 | ||
| 1750 | subsection \<open>Multi-threading with shared memory\<close> | |
| 1751 | ||
| 61854 | 1752 | text \<open> | 
| 1753 | Multiple threads help to organize advanced operations of the system, such as | |
| 1754 | real-time conditions on command transactions, sub-components with explicit | |
| 1755 | communication, general asynchronous interaction etc. Moreover, parallel | |
| 1756 | evaluation is a prerequisite to make adequate use of the CPU resources that | |
| 1757 | are available on multi-core systems.\<^footnote>\<open>Multi-core computing does not mean | |
| 1758 | that there are ``spare cycles'' to be wasted. It means that the continued | |
| 1759 | exponential speedup of CPU performance due to ``Moore's Law'' follows | |
| 1760 | different rules: clock frequency has reached its peak around 2005, and | |
| 1761 | applications need to be parallelized in order to avoid a perceived loss of | |
| 76987 | 1762 | performance. See also \<^cite>\<open>"Sutter:2005"\<close>.\<close> | 
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1763 | |
| 57421 | 1764 | Isabelle/Isar exploits the inherent structure of theories and proofs to | 
| 61854 | 1765 | support \<^emph>\<open>implicit parallelism\<close> to a large extent. LCF-style theorem proving | 
| 76987 | 1766 | provides almost ideal conditions for that, see also \<^cite>\<open>"Wenzel:2009"\<close>. | 
| 61854 | 1767 | This means, significant parts of theory and proof checking is parallelized | 
| 1768 | by default. In Isabelle2013, a maximum speedup-factor of 3.5 on 4 cores and | |
| 76987 | 1769 | 6.5 on 8 cores can be expected \<^cite>\<open>"Wenzel:2013:ITP"\<close>. | 
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1770 | |
| 61416 | 1771 | \<^medskip> | 
| 61854 | 1772 | ML threads lack the memory protection of separate processes, and operate | 
| 1773 | concurrently on shared heap memory. This has the advantage that results of | |
| 1774 | independent computations are directly available to other threads: abstract | |
| 1775 | values can be passed without copying or awkward serialization that is | |
| 1776 | typically required for separate processes. | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1777 | |
| 61854 | 1778 | To make shared-memory multi-threading work robustly and efficiently, some | 
| 1779 | programming guidelines need to be observed. While the ML system is | |
| 1780 | responsible to maintain basic integrity of the representation of ML values | |
| 1781 | in memory, the application programmer needs to ensure that multi-threaded | |
| 1782 | execution does not break the intended semantics. | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1783 | |
| 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1784 |   \begin{warn}
 | 
| 61854 | 1785 | To participate in implicit parallelism, tools need to be thread-safe. A | 
| 1786 | single ill-behaved tool can affect the stability and performance of the | |
| 1787 | whole system. | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1788 |   \end{warn}
 | 
| 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1789 | |
| 57421 | 1790 | Apart from observing the principles of thread-safeness passively, advanced | 
| 1791 | tools may also exploit parallelism actively, e.g.\ by using library | |
| 39868 | 1792 |   functions for parallel list operations (\secref{sec:parlist}).
 | 
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1793 | |
| 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1794 |   \begin{warn}
 | 
| 61854 | 1795 | Parallel computing resources are managed centrally by the Isabelle/ML | 
| 1796 | infrastructure. User programs should not fork their own ML threads to | |
| 1797 | perform heavy computations. | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1798 |   \end{warn}
 | 
| 58618 | 1799 | \<close> | 
| 1800 | ||
| 1801 | ||
| 1802 | subsection \<open>Critical shared resources\<close> | |
| 1803 | ||
| 61854 | 1804 | text \<open> | 
| 1805 | Thread-safeness is mainly concerned about concurrent read/write access to | |
| 1806 | shared resources, which are outside the purely functional world of ML. This | |
| 1807 | covers the following in particular. | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1808 | |
| 61854 | 1809 | \<^item> Global references (or arrays), i.e.\ mutable memory cells that persist | 
| 1810 | over several invocations of associated operations.\<^footnote>\<open>This is independent of | |
| 1811 | the visibility of such mutable values in the toplevel scope.\<close> | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1812 | |
| 61854 | 1813 | \<^item> Global state of the running Isabelle/ML process, i.e.\ raw I/O channels, | 
| 1814 | environment variables, current working directory. | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1815 | |
| 61854 | 1816 | \<^item> Writable resources in the file-system that are shared among different | 
| 1817 | threads or external processes. | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1818 | |
| 61854 | 1819 | Isabelle/ML provides various mechanisms to avoid critical shared resources | 
| 1820 | in most situations. As last resort there are some mechanisms for explicit | |
| 1821 | synchronization. The following guidelines help to make Isabelle/ML programs | |
| 1822 | work smoothly in a concurrent environment. | |
| 1823 | ||
| 1824 | \<^item> Avoid global references altogether. Isabelle/Isar maintains a uniform | |
| 1825 | context that incorporates arbitrary data declared by user programs | |
| 1826 |   (\secref{sec:context-data}). This context is passed as plain value and user
 | |
| 1827 | tools can get/map their own data in a purely functional manner. | |
| 1828 |   Configuration options within the context (\secref{sec:config-options})
 | |
| 1829 | provide simple drop-in replacements for historic reference variables. | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1830 | |
| 61854 | 1831 | \<^item> Keep components with local state information re-entrant. Instead of poking | 
| 1832 | initial values into (private) global references, a new state record can be | |
| 1833 | created on each invocation, and passed through any auxiliary functions of | |
| 1834 | the component. The state record contain mutable references in special | |
| 1835 | situations, without requiring any synchronization, as long as each | |
| 1836 | invocation gets its own copy and the tool itself is single-threaded. | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1837 | |
| 61854 | 1838 | \<^item> Avoid raw output on \<open>stdout\<close> or \<open>stderr\<close>. The Poly/ML library is | 
| 1839 | thread-safe for each individual output operation, but the ordering of | |
| 1840 | parallel invocations is arbitrary. This means raw output will appear on some | |
| 1841 | system console with unpredictable interleaving of atomic chunks. | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1842 | |
| 39868 | 1843 | Note that this does not affect regular message output channels | 
| 61854 | 1844 |   (\secref{sec:message-channels}). An official message id is associated with
 | 
| 1845 | the command transaction from where it originates, independently of other | |
| 1846 | transactions. This means each running Isar command has effectively its own | |
| 1847 | set of message channels, and interleaving can only happen when commands use | |
| 1848 | parallelism internally (and only at message boundaries). | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1849 | |
| 61854 | 1850 | \<^item> Treat environment variables and the current working directory of the | 
| 1851 | running process as read-only. | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1852 | |
| 61854 | 1853 | \<^item> Restrict writing to the file-system to unique temporary files. Isabelle | 
| 1854 | already provides a temporary directory that is unique for the running | |
| 1855 | process, and there is a centralized source of unique serial numbers in | |
| 1856 | Isabelle/ML. Thus temporary files that are passed to to some external | |
| 1857 | process will be always disjoint, and thus thread-safe. | |
| 58618 | 1858 | \<close> | 
| 1859 | ||
| 1860 | text %mlref \<open> | |
| 39868 | 1861 |   \begin{mldecls}
 | 
| 73764 | 1862 |   @{define_ML File.tmp_path: "Path.T -> Path.T"} \\
 | 
| 1863 |   @{define_ML serial_string: "unit -> string"} \\
 | |
| 39868 | 1864 |   \end{mldecls}
 | 
| 1865 | ||
| 69597 | 1866 | \<^descr> \<^ML>\<open>File.tmp_path\<close>~\<open>path\<close> relocates the base component of \<open>path\<close> into the | 
| 61854 | 1867 | unique temporary directory of the running Isabelle/ML process. | 
| 39868 | 1868 | |
| 69597 | 1869 | \<^descr> \<^ML>\<open>serial_string\<close>~\<open>()\<close> creates a new serial number that is unique over | 
| 61854 | 1870 | the runtime of the Isabelle/ML process. | 
| 58618 | 1871 | \<close> | 
| 1872 | ||
| 61854 | 1873 | text %mlex \<open> | 
| 1874 | The following example shows how to create unique temporary file names. | |
| 58618 | 1875 | \<close> | 
| 1876 | ||
| 59902 | 1877 | ML_val \<open> | 
| 39868 | 1878 |   val tmp1 = File.tmp_path (Path.basic ("foo" ^ serial_string ()));
 | 
| 1879 |   val tmp2 = File.tmp_path (Path.basic ("foo" ^ serial_string ()));
 | |
| 69597 | 1880 | \<^assert> (tmp1 <> tmp2); | 
| 58618 | 1881 | \<close> | 
| 1882 | ||
| 1883 | ||
| 1884 | subsection \<open>Explicit synchronization\<close> | |
| 1885 | ||
| 61854 | 1886 | text \<open> | 
| 1887 | Isabelle/ML provides explicit synchronization for mutable variables over | |
| 59180 
c0fa3b3bdabd
discontinued central critical sections: NAMED_CRITICAL / CRITICAL;
 wenzelm parents: 
59138diff
changeset | 1888 | immutable data, which may be updated atomically and exclusively. This | 
| 
c0fa3b3bdabd
discontinued central critical sections: NAMED_CRITICAL / CRITICAL;
 wenzelm parents: 
59138diff
changeset | 1889 | addresses the rare situations where mutable shared resources are really | 
| 
c0fa3b3bdabd
discontinued central critical sections: NAMED_CRITICAL / CRITICAL;
 wenzelm parents: 
59138diff
changeset | 1890 | required. Synchronization in Isabelle/ML is based on primitives of Poly/ML, | 
| 
c0fa3b3bdabd
discontinued central critical sections: NAMED_CRITICAL / CRITICAL;
 wenzelm parents: 
59138diff
changeset | 1891 | which have been adapted to the specific assumptions of the concurrent | 
| 
c0fa3b3bdabd
discontinued central critical sections: NAMED_CRITICAL / CRITICAL;
 wenzelm parents: 
59138diff
changeset | 1892 | Isabelle environment. User code should not break this abstraction, but stay | 
| 
c0fa3b3bdabd
discontinued central critical sections: NAMED_CRITICAL / CRITICAL;
 wenzelm parents: 
59138diff
changeset | 1893 | within the confines of concurrent Isabelle/ML. | 
| 
c0fa3b3bdabd
discontinued central critical sections: NAMED_CRITICAL / CRITICAL;
 wenzelm parents: 
59138diff
changeset | 1894 | |
| 61854 | 1895 | A \<^emph>\<open>synchronized variable\<close> is an explicit state component associated with | 
| 1896 | mechanisms for locking and signaling. There are operations to await a | |
| 59180 
c0fa3b3bdabd
discontinued central critical sections: NAMED_CRITICAL / CRITICAL;
 wenzelm parents: 
59138diff
changeset | 1897 | condition, change the state, and signal the change to all other waiting | 
| 61477 | 1898 | threads. Synchronized access to the state variable is \<^emph>\<open>not\<close> re-entrant: | 
| 61854 | 1899 | direct or indirect nesting within the same thread will cause a deadlock! | 
| 1900 | \<close> | |
| 58618 | 1901 | |
| 1902 | text %mlref \<open> | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1903 |   \begin{mldecls}
 | 
| 73769 | 1904 |   @{define_ML_type 'a "Synchronized.var"} \\
 | 
| 73764 | 1905 |   @{define_ML Synchronized.var: "string -> 'a -> 'a Synchronized.var"} \\
 | 
| 1906 |   @{define_ML Synchronized.guarded_access: "'a Synchronized.var ->
 | |
| 39871 | 1907 |   ('a -> ('b * 'a) option) -> 'b"} \\
 | 
| 1908 |   \end{mldecls}
 | |
| 39867 
a8363532cd4d
somewhat modernized version of "Thread-safe programming";
 wenzelm parents: 
39866diff
changeset | 1909 | |
| 69597 | 1910 | \<^descr> Type \<^ML_type>\<open>'a Synchronized.var\<close> represents synchronized variables | 
| 1911 | with state of type \<^ML_type>\<open>'a\<close>. | |
| 1912 | ||
| 1913 | \<^descr> \<^ML>\<open>Synchronized.var\<close>~\<open>name x\<close> creates a synchronized variable that is | |
| 61854 | 1914 | initialized with value \<open>x\<close>. The \<open>name\<close> is used for tracing. | 
| 61493 | 1915 | |
| 69597 | 1916 | \<^descr> \<^ML>\<open>Synchronized.guarded_access\<close>~\<open>var f\<close> lets the function \<open>f\<close> operate | 
| 61854 | 1917 | within a critical section on the state \<open>x\<close> as follows: if \<open>f x\<close> produces | 
| 69597 | 1918 | \<^ML>\<open>NONE\<close>, it continues to wait on the internal condition variable, | 
| 61854 | 1919 | expecting that some other thread will eventually change the content in a | 
| 69597 | 1920 | suitable manner; if \<open>f x\<close> produces \<^ML>\<open>SOME\<close>~\<open>(y, x')\<close> it is satisfied and | 
| 61854 | 1921 | assigns the new state value \<open>x'\<close>, broadcasts a signal to all waiting threads | 
| 1922 | on the associated condition variable, and returns the result \<open>y\<close>. | |
| 39871 | 1923 | |
| 69597 | 1924 | There are some further variants of the \<^ML>\<open>Synchronized.guarded_access\<close> | 
| 63680 | 1925 | combinator, see \<^file>\<open>~~/src/Pure/Concurrent/synchronized.ML\<close> for details. | 
| 58618 | 1926 | \<close> | 
| 1927 | ||
| 61854 | 1928 | text %mlex \<open> | 
| 1929 | The following example implements a counter that produces positive integers | |
| 1930 | that are unique over the runtime of the Isabelle process: | |
| 1931 | \<close> | |
| 58618 | 1932 | |
| 59902 | 1933 | ML_val \<open> | 
| 39871 | 1934 | local | 
| 1935 | val counter = Synchronized.var "counter" 0; | |
| 1936 | in | |
| 1937 | fun next () = | |
| 1938 | Synchronized.guarded_access counter | |
| 1939 | (fn i => | |
| 1940 | let val j = i + 1 | |
| 1941 | in SOME (j, j) end); | |
| 1942 | end; | |
| 59902 | 1943 | |
| 39871 | 1944 | val a = next (); | 
| 1945 | val b = next (); | |
| 69597 | 1946 | \<^assert> (a <> b); | 
| 58618 | 1947 | \<close> | 
| 1948 | ||
| 61416 | 1949 | text \<open> | 
| 1950 | \<^medskip> | |
| 63680 | 1951 | See \<^file>\<open>~~/src/Pure/Concurrent/mailbox.ML\<close> how to implement a mailbox as | 
| 1952 | synchronized variable over a purely functional list. | |
| 61854 | 1953 | \<close> | 
| 58618 | 1954 | |
| 1955 | ||
| 78742 | 1956 | section \<open>Managed evaluation \label{sec:managed-eval}\<close>
 | 
| 58618 | 1957 | |
| 61854 | 1958 | text \<open> | 
| 1959 | Execution of Standard ML follows the model of strict functional evaluation | |
| 1960 | with optional exceptions. Evaluation happens whenever some function is | |
| 1961 | applied to (sufficiently many) arguments. The result is either an explicit | |
| 1962 | value or an implicit exception. | |
| 52419 | 1963 | |
| 61854 | 1964 | \<^emph>\<open>Managed evaluation\<close> in Isabelle/ML organizes expressions and results to | 
| 1965 | control certain physical side-conditions, to say more specifically when and | |
| 1966 | how evaluation happens. For example, the Isabelle/ML library supports lazy | |
| 1967 | evaluation with memoing, parallel evaluation via futures, asynchronous | |
| 1968 | evaluation via promises, evaluation with time limit etc. | |
| 52419 | 1969 | |
| 61416 | 1970 | \<^medskip> | 
| 66663 | 1971 | An \<^emph>\<open>unevaluated expression\<close> is represented either as unit abstraction | 
| 1972 | \<^verbatim>\<open>fn () => a\<close> of type \<^verbatim>\<open>unit -> 'a\<close> or as regular function \<^verbatim>\<open>fn a => b\<close> of | |
| 1973 | type \<^verbatim>\<open>'a -> 'b\<close>. Both forms occur routinely, and special care is required | |
| 1974 | to tell them apart --- the static type-system of SML is only of limited help | |
| 61854 | 1975 | here. | 
| 52419 | 1976 | |
| 66663 | 1977 | The first form is more intuitive: some combinator \<^verbatim>\<open>(unit -> 'a) -> 'a\<close> | 
| 1978 | applies the given function to \<^verbatim>\<open>()\<close> to initiate the postponed evaluation | |
| 1979 | process. The second form is more flexible: some combinator | |
| 1980 |   \<^verbatim>\<open>('a -> 'b) -> 'a -> 'b\<close> acts like a modified form of function application;
 | |
| 1981 | several such combinators may be cascaded to modify a given function, before | |
| 1982 | it is ultimately applied to some argument. | |
| 52419 | 1983 | |
| 61416 | 1984 | \<^medskip> | 
| 61854 | 1985 | \<^emph>\<open>Reified results\<close> make the disjoint sum of regular values versions | 
| 1986 | exceptional situations explicit as ML datatype: \<open>'a result = Res of 'a | Exn | |
| 1987 | of exn\<close>. This is typically used for administrative purposes, to store the | |
| 1988 | overall outcome of an evaluation process. | |
| 52419 | 1989 | |
| 61854 | 1990 | \<^emph>\<open>Parallel exceptions\<close> aggregate reified results, such that multiple | 
| 1991 | exceptions are digested as a collection in canonical form that identifies | |
| 1992 | exceptions according to their original occurrence. This is particular | |
| 1993 |   important for parallel evaluation via futures \secref{sec:futures}, which
 | |
| 1994 | are organized as acyclic graph of evaluations that depend on other | |
| 1995 | evaluations: exceptions stemming from shared sub-graphs are exposed exactly | |
| 1996 | once and in the order of their original occurrence (e.g.\ when printed at | |
| 1997 | the toplevel). Interrupt counts as neutral element here: it is treated as | |
| 1998 | minimal information about some canceled evaluation process, and is absorbed | |
| 1999 | by the presence of regular program exceptions. | |
| 2000 | \<close> | |
| 58618 | 2001 | |
| 2002 | text %mlref \<open> | |
| 52419 | 2003 |   \begin{mldecls}
 | 
| 73769 | 2004 |   @{define_ML_type 'a "Exn.result"} \\
 | 
| 73764 | 2005 |   @{define_ML Exn.capture: "('a -> 'b) -> 'a -> 'b Exn.result"} \\
 | 
| 78705 
fde0b195cb7d
clarified signature: avoid association with potentially dangerous Exn.capture;
 wenzelm parents: 
78681diff
changeset | 2006 |   @{define_ML Exn.result: "('a -> 'b) -> 'a -> 'b Exn.result"} \\
 | 
| 73764 | 2007 |   @{define_ML Exn.release: "'a Exn.result -> 'a"} \\
 | 
| 2008 |   @{define_ML Par_Exn.release_all: "'a Exn.result list -> 'a list"} \\
 | |
| 2009 |   @{define_ML Par_Exn.release_first: "'a Exn.result list -> 'a list"} \\
 | |
| 52419 | 2010 |   \end{mldecls}
 | 
| 2011 | ||
| 69597 | 2012 | \<^descr> Type \<^ML_type>\<open>'a Exn.result\<close> represents the disjoint sum of ML results | 
| 2013 | explicitly, with constructor \<^ML>\<open>Exn.Res\<close> for regular values and \<^ML>\<open>Exn.Exn\<close> for exceptions. | |
| 2014 | ||
| 2015 | \<^descr> \<^ML>\<open>Exn.capture\<close>~\<open>f x\<close> manages the evaluation of \<open>f x\<close> such that | |
| 2016 | exceptions are made explicit as \<^ML>\<open>Exn.Exn\<close>. Note that this includes | |
| 61854 | 2017 |   physical interrupts (see also \secref{sec:exceptions}), so the same
 | 
| 2018 | precautions apply to user code: interrupts must not be absorbed | |
| 2019 | accidentally! | |
| 52419 | 2020 | |
| 78705 
fde0b195cb7d
clarified signature: avoid association with potentially dangerous Exn.capture;
 wenzelm parents: 
78681diff
changeset | 2021 | \<^descr> \<^ML>\<open>Exn.result\<close> is similar to \<^ML>\<open>Exn.capture\<close>, but interrupts are | 
| 
fde0b195cb7d
clarified signature: avoid association with potentially dangerous Exn.capture;
 wenzelm parents: 
78681diff
changeset | 2022 | immediately re-raised as required for user code. | 
| 52419 | 2023 | |
| 69597 | 2024 | \<^descr> \<^ML>\<open>Exn.release\<close>~\<open>result\<close> releases the original runtime result, exposing | 
| 61854 | 2025 | its regular value or raising the reified exception. | 
| 52419 | 2026 | |
| 69597 | 2027 | \<^descr> \<^ML>\<open>Par_Exn.release_all\<close>~\<open>results\<close> combines results that were produced | 
| 61854 | 2028 | independently (e.g.\ by parallel evaluation). If all results are regular | 
| 2029 | values, that list is returned. Otherwise, the collection of all exceptions | |
| 2030 | is raised, wrapped-up as collective parallel exception. Note that the latter | |
| 2031 | prevents access to individual exceptions by conventional \<^verbatim>\<open>handle\<close> of ML. | |
| 52419 | 2032 | |
| 69597 | 2033 | \<^descr> \<^ML>\<open>Par_Exn.release_first\<close> is similar to \<^ML>\<open>Par_Exn.release_all\<close>, but | 
| 61854 | 2034 | only the first (meaningful) exception that has occurred in the original | 
| 2035 | evaluation process is raised again, the others are ignored. That single | |
| 2036 | exception may get handled by conventional means in ML. | |
| 58618 | 2037 | \<close> | 
| 2038 | ||
| 2039 | ||
| 2040 | subsection \<open>Parallel skeletons \label{sec:parlist}\<close>
 | |
| 2041 | ||
| 2042 | text \<open> | |
| 61854 | 2043 | Algorithmic skeletons are combinators that operate on lists in parallel, in | 
| 2044 | the manner of well-known \<open>map\<close>, \<open>exists\<close>, \<open>forall\<close> etc. Management of | |
| 2045 |   futures (\secref{sec:futures}) and their results as reified exceptions is
 | |
| 2046 | wrapped up into simple programming interfaces that resemble the sequential | |
| 2047 | versions. | |
| 52420 | 2048 | |
| 61854 | 2049 | What remains is the application-specific problem to present expressions with | 
| 2050 | suitable \<^emph>\<open>granularity\<close>: each list element corresponds to one evaluation | |
| 2051 | task. If the granularity is too coarse, the available CPUs are not | |
| 2052 | saturated. If it is too fine-grained, CPU cycles are wasted due to the | |
| 2053 | overhead of organizing parallel processing. In the worst case, parallel | |
| 52420 | 2054 | performance will be less than the sequential counterpart! | 
| 58618 | 2055 | \<close> | 
| 2056 | ||
| 2057 | text %mlref \<open> | |
| 52420 | 2058 |   \begin{mldecls}
 | 
| 73764 | 2059 |   @{define_ML Par_List.map: "('a -> 'b) -> 'a list -> 'b list"} \\
 | 
| 2060 |   @{define_ML Par_List.get_some: "('a -> 'b option) -> 'a list -> 'b option"} \\
 | |
| 52420 | 2061 |   \end{mldecls}
 | 
| 2062 | ||
| 69597 | 2063 | \<^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 | 2064 | x\<^sub>n]\<close>, but the evaluation of \<open>f x\<^sub>i\<close> for \<open>i = 1, \<dots>, n\<close> is performed in | 
| 2065 | parallel. | |
| 61493 | 2066 | |
| 61854 | 2067 | An exception in any \<open>f x\<^sub>i\<close> cancels the overall evaluation process. The | 
| 69597 | 2068 | final result is produced via \<^ML>\<open>Par_Exn.release_first\<close> as explained above, | 
| 61854 | 2069 | which means the first program exception that happened to occur in the | 
| 2070 | parallel evaluation is propagated, and all other failures are ignored. | |
| 52420 | 2071 | |
| 69597 | 2072 | \<^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 | 2073 | the form \<open>SOME y\<^sub>i\<close>, if that exists, otherwise \<open>NONE\<close>. Thus it is similar to | 
| 69597 | 2074 | \<^ML>\<open>Library.get_first\<close>, but subject to a non-deterministic parallel choice | 
| 61854 | 2075 | process. The first successful result cancels the overall evaluation process; | 
| 69597 | 2076 | other exceptions are propagated as for \<^ML>\<open>Par_List.map\<close>. | 
| 52420 | 2077 | |
| 61854 | 2078 | This generic parallel choice combinator is the basis for derived forms, such | 
| 69597 | 2079 | as \<^ML>\<open>Par_List.find_some\<close>, \<^ML>\<open>Par_List.exists\<close>, \<^ML>\<open>Par_List.forall\<close>. | 
| 58618 | 2080 | \<close> | 
| 2081 | ||
| 61854 | 2082 | text %mlex \<open> | 
| 2083 | Subsequently, the Ackermann function is evaluated in parallel for some | |
| 2084 | ranges of arguments. | |
| 2085 | \<close> | |
| 58618 | 2086 | |
| 2087 | ML_val \<open> | |
| 52420 | 2088 | fun ackermann 0 n = n + 1 | 
| 2089 | | ackermann m 0 = ackermann (m - 1) 1 | |
| 2090 | | ackermann m n = ackermann (m - 1) (ackermann m (n - 1)); | |
| 2091 | ||
| 2092 | Par_List.map (ackermann 2) (500 upto 1000); | |
| 2093 | Par_List.map (ackermann 3) (5 upto 10); | |
| 58618 | 2094 | \<close> | 
| 2095 | ||
| 2096 | ||
| 2097 | subsection \<open>Lazy evaluation\<close> | |
| 2098 | ||
| 2099 | text \<open> | |
| 61854 | 2100 | Classic lazy evaluation works via the \<open>lazy\<close>~/ \<open>force\<close> pair of operations: | 
| 2101 | \<open>lazy\<close> to wrap an unevaluated expression, and \<open>force\<close> to evaluate it once | |
| 2102 | and store its result persistently. Later invocations of \<open>force\<close> retrieve the | |
| 2103 | stored result without another evaluation. Isabelle/ML refines this idea to | |
| 2104 | accommodate the aspects of multi-threading, synchronous program exceptions | |
| 2105 | and asynchronous interrupts. | |
| 57347 | 2106 | |
| 61854 | 2107 | The first thread that invokes \<open>force\<close> on an unfinished lazy value changes | 
| 2108 | its state into a \<^emph>\<open>promise\<close> of the eventual result and starts evaluating it. | |
| 2109 | Any other threads that \<open>force\<close> the same lazy value in the meantime need to | |
| 2110 | wait for it to finish, by producing a regular result or program exception. | |
| 2111 | If the evaluation attempt is interrupted, this event is propagated to all | |
| 2112 | waiting threads and the lazy value is reset to its original state. | |
| 57347 | 2113 | |
| 2114 | This means a lazy value is completely evaluated at most once, in a | |
| 2115 | thread-safe manner. There might be multiple interrupted evaluation attempts, | |
| 2116 | and multiple receivers of intermediate interrupt events. Interrupts are | |
| 61477 | 2117 | \<^emph>\<open>not\<close> made persistent: later evaluation attempts start again from the | 
| 57347 | 2118 | original expression. | 
| 58618 | 2119 | \<close> | 
| 2120 | ||
| 2121 | text %mlref \<open> | |
| 57347 | 2122 |   \begin{mldecls}
 | 
| 73769 | 2123 |   @{define_ML_type 'a "lazy"} \\
 | 
| 73764 | 2124 |   @{define_ML Lazy.lazy: "(unit -> 'a) -> 'a lazy"} \\
 | 
| 2125 |   @{define_ML Lazy.value: "'a -> 'a lazy"} \\
 | |
| 2126 |   @{define_ML Lazy.force: "'a lazy -> 'a"} \\
 | |
| 57347 | 2127 |   \end{mldecls}
 | 
| 2128 | ||
| 69597 | 2129 | \<^descr> Type \<^ML_type>\<open>'a lazy\<close> represents lazy values over type \<^verbatim>\<open>'a\<close>. | 
| 2130 | ||
| 2131 | \<^descr> \<^ML>\<open>Lazy.lazy\<close>~\<open>(fn () => e)\<close> wraps the unevaluated expression \<open>e\<close> as | |
| 61854 | 2132 | unfinished lazy value. | 
| 61493 | 2133 | |
| 69597 | 2134 | \<^descr> \<^ML>\<open>Lazy.value\<close>~\<open>a\<close> wraps the value \<open>a\<close> as finished lazy value. When | 
| 61854 | 2135 | forced, it returns \<open>a\<close> without any further evaluation. | 
| 57347 | 2136 | |
| 57349 | 2137 | There is very low overhead for this proforma wrapping of strict values as | 
| 2138 | lazy values. | |
| 57347 | 2139 | |
| 69597 | 2140 | \<^descr> \<^ML>\<open>Lazy.force\<close>~\<open>x\<close> produces the result of the lazy value in a | 
| 57347 | 2141 | thread-safe manner as explained above. Thus it may cause the current thread | 
| 2142 | to wait on a pending evaluation attempt by another thread. | |
| 58618 | 2143 | \<close> | 
| 2144 | ||
| 2145 | ||
| 2146 | subsection \<open>Futures \label{sec:futures}\<close>
 | |
| 2147 | ||
| 2148 | text \<open> | |
| 57349 | 2149 | Futures help to organize parallel execution in a value-oriented manner, with | 
| 61854 | 2150 | \<open>fork\<close>~/ \<open>join\<close> as the main pair of operations, and some further variants; | 
| 76987 | 2151 | see also \<^cite>\<open>"Wenzel:2009" and "Wenzel:2013:ITP"\<close>. Unlike lazy values, | 
| 61854 | 2152 | futures are evaluated strictly and spontaneously on separate worker threads. | 
| 2153 | Futures may be canceled, which leads to interrupts on running evaluation | |
| 2154 | attempts, and forces structurally related futures to fail for all time; | |
| 2155 | already finished futures remain unchanged. Exceptions between related | |
| 57350 | 2156 | futures are propagated as well, and turned into parallel exceptions (see | 
| 2157 | above). | |
| 57349 | 2158 | |
| 2159 | Technically, a future is a single-assignment variable together with a | |
| 61854 | 2160 | \<^emph>\<open>task\<close> that serves administrative purposes, notably within the \<^emph>\<open>task | 
| 2161 | queue\<close> where new futures are registered for eventual evaluation and the | |
| 2162 | worker threads retrieve their work. | |
| 57349 | 2163 | |
| 57350 | 2164 | The pool of worker threads is limited, in correlation with the number of | 
| 2165 | physical cores on the machine. Note that allocation of runtime resources may | |
| 2166 | be distorted either if workers yield CPU time (e.g.\ via system sleep or | |
| 2167 | wait operations), or if non-worker threads contend for significant runtime | |
| 2168 | resources independently. There is a limited number of replacement worker | |
| 2169 | threads that get activated in certain explicit wait conditions, after a | |
| 2170 | timeout. | |
| 2171 | ||
| 61416 | 2172 | \<^medskip> | 
| 61854 | 2173 | Each future task belongs to some \<^emph>\<open>task group\<close>, which represents the | 
| 2174 | hierarchic structure of related tasks, together with the exception status a | |
| 2175 | that point. By default, the task group of a newly created future is a new | |
| 2176 | sub-group of the presently running one, but it is also possible to indicate | |
| 2177 | different group layouts under program control. | |
| 57349 | 2178 | |
| 2179 | Cancellation of futures actually refers to the corresponding task group and | |
| 2180 | all its sub-groups. Thus interrupts are propagated down the group hierarchy. | |
| 2181 | Regular program exceptions are treated likewise: failure of the evaluation | |
| 2182 | of some future task affects its own group and all sub-groups. Given a | |
| 61854 | 2183 | particular task group, its \<^emph>\<open>group status\<close> cumulates all relevant exceptions | 
| 2184 | according to its position within the group hierarchy. Interrupted tasks that | |
| 2185 | lack regular result information, will pick up parallel exceptions from the | |
| 2186 | cumulative group status. | |
| 57349 | 2187 | |
| 61416 | 2188 | \<^medskip> | 
| 61854 | 2189 | A \<^emph>\<open>passive future\<close> or \<^emph>\<open>promise\<close> is a future with slightly different | 
| 2190 | evaluation policies: there is only a single-assignment variable and some | |
| 2191 | expression to evaluate for the \<^emph>\<open>failed\<close> case (e.g.\ to clean up resources | |
| 2192 | when canceled). A regular result is produced by external means, using a | |
| 2193 | separate \<^emph>\<open>fulfill\<close> operation. | |
| 57349 | 2194 | |
| 2195 | Promises are managed in the same task queue, so regular futures may depend | |
| 2196 | on them. This allows a form of reactive programming, where some promises are | |
| 2197 | used as minimal elements (or guards) within the future dependency graph: | |
| 2198 | when these promises are fulfilled the evaluation of subsequent futures | |
| 2199 | starts spontaneously, according to their own inter-dependencies. | |
| 58618 | 2200 | \<close> | 
| 2201 | ||
| 2202 | text %mlref \<open> | |
| 57348 | 2203 |   \begin{mldecls}
 | 
| 73769 | 2204 |   @{define_ML_type 'a "future"} \\
 | 
| 73764 | 2205 |   @{define_ML Future.fork: "(unit -> 'a) -> 'a future"} \\
 | 
| 2206 |   @{define_ML Future.forks: "Future.params -> (unit -> 'a) list -> 'a future list"} \\
 | |
| 2207 |   @{define_ML Future.join: "'a future -> 'a"} \\
 | |
| 2208 |   @{define_ML Future.joins: "'a future list -> 'a list"} \\
 | |
| 2209 |   @{define_ML Future.value: "'a -> 'a future"} \\
 | |
| 2210 |   @{define_ML Future.map: "('a -> 'b) -> 'a future -> 'b future"} \\
 | |
| 2211 |   @{define_ML Future.cancel: "'a future -> unit"} \\
 | |
| 2212 |   @{define_ML Future.cancel_group: "Future.group -> unit"} \\[0.5ex]
 | |
| 2213 |   @{define_ML Future.promise: "(unit -> unit) -> 'a future"} \\
 | |
| 2214 |   @{define_ML Future.fulfill: "'a future -> 'a -> unit"} \\
 | |
| 57348 | 2215 |   \end{mldecls}
 | 
| 2216 | ||
| 69597 | 2217 | \<^descr> Type \<^ML_type>\<open>'a future\<close> represents future values over type \<^verbatim>\<open>'a\<close>. | 
| 2218 | ||
| 2219 | \<^descr> \<^ML>\<open>Future.fork\<close>~\<open>(fn () => e)\<close> registers the unevaluated expression \<open>e\<close> | |
| 61854 | 2220 | as unfinished future value, to be evaluated eventually on the parallel | 
| 69597 | 2221 | worker-thread farm. This is a shorthand for \<^ML>\<open>Future.forks\<close> below, with | 
| 61854 | 2222 | default parameters and a single expression. | 
| 57348 | 2223 | |
| 69597 | 2224 | \<^descr> \<^ML>\<open>Future.forks\<close>~\<open>params exprs\<close> is the general interface to fork several | 
| 61854 | 2225 | futures simultaneously. The \<open>params\<close> consist of the following fields: | 
| 57348 | 2226 | |
| 69597 | 2227 | \<^item> \<open>name : string\<close> (default \<^ML>\<open>""\<close>) specifies a common name for the | 
| 61854 | 2228 | tasks of the forked futures, which serves diagnostic purposes. | 
| 61458 | 2229 | |
| 69597 | 2230 | \<^item> \<open>group : Future.group option\<close> (default \<^ML>\<open>NONE\<close>) specifies an optional | 
| 2231 | task group for the forked futures. \<^ML>\<open>NONE\<close> means that a new sub-group | |
| 61854 | 2232 | of the current worker-thread task context is created. If this is not a | 
| 2233 | worker thread, the group will be a new root in the group hierarchy. | |
| 61458 | 2234 | |
| 69597 | 2235 | \<^item> \<open>deps : Future.task list\<close> (default \<^ML>\<open>[]\<close>) specifies dependencies on | 
| 61854 | 2236 | other future tasks, i.e.\ the adjacency relation in the global task queue. | 
| 2237 | Dependencies on already finished tasks are ignored. | |
| 61458 | 2238 | |
| 69597 | 2239 | \<^item> \<open>pri : int\<close> (default \<^ML>\<open>0\<close>) specifies a priority within the task | 
| 61854 | 2240 | queue. | 
| 61458 | 2241 | |
| 75617 | 2242 | Typically there is only little deviation from the default priority | 
| 2243 | \<^ML>\<open>0\<close>. As a rule of thumb, \<^ML>\<open>~1\<close> means ``low priority" and | |
| 2244 | \<^ML>\<open>1\<close> means ``high priority''. | |
| 61458 | 2245 | |
| 61854 | 2246 | Note that the task priority only affects the position in the queue, not | 
| 2247 | the thread priority. When a worker thread picks up a task for processing, | |
| 2248 | it runs with the normal thread priority to the end (or until canceled). | |
| 2249 | Higher priority tasks that are queued later need to wait until this (or | |
| 2250 | another) worker thread becomes free again. | |
| 61458 | 2251 | |
| 69597 | 2252 | \<^item> \<open>interrupts : bool\<close> (default \<^ML>\<open>true\<close>) tells whether the worker thread | 
| 61854 | 2253 | that processes the corresponding task is initially put into interruptible | 
| 2254 | state. This state may change again while running, by modifying the thread | |
| 2255 | attributes. | |
| 61458 | 2256 | |
| 61854 | 2257 | With interrupts disabled, a running future task cannot be canceled. It is | 
| 61458 | 2258 | the responsibility of the programmer that this special state is retained | 
| 2259 | only briefly. | |
| 57348 | 2260 | |
| 69597 | 2261 | \<^descr> \<^ML>\<open>Future.join\<close>~\<open>x\<close> retrieves the value of an already finished future, | 
| 61854 | 2262 | which may lead to an exception, according to the result of its previous | 
| 2263 | evaluation. | |
| 57348 | 2264 | |
| 2265 | For an unfinished future there are several cases depending on the role of | |
| 2266 | the current thread and the status of the future. A non-worker thread waits | |
| 2267 | passively until the future is eventually evaluated. A worker thread | |
| 2268 | temporarily changes its task context and takes over the responsibility to | |
| 57349 | 2269 | evaluate the future expression on the spot. The latter is done in a | 
| 2270 | thread-safe manner: other threads that intend to join the same future need | |
| 2271 | to wait until the ongoing evaluation is finished. | |
| 2272 | ||
| 2273 | Note that excessive use of dynamic dependencies of futures by adhoc joining | |
| 2274 | may lead to bad utilization of CPU cores, due to threads waiting on other | |
| 2275 | threads to finish required futures. The future task farm has a limited | |
| 2276 | amount of replacement threads that continue working on unrelated tasks after | |
| 2277 | some timeout. | |
| 57348 | 2278 | |
| 2279 | Whenever possible, static dependencies of futures should be specified | |
| 61854 | 2280 | explicitly when forked (see \<open>deps\<close> above). Thus the evaluation can work from | 
| 2281 | the bottom up, without join conflicts and wait states. | |
| 57349 | 2282 | |
| 69597 | 2283 | \<^descr> \<^ML>\<open>Future.joins\<close>~\<open>xs\<close> joins the given list of futures simultaneously, | 
| 2284 | which is more efficient than \<^ML>\<open>map Future.join\<close>~\<open>xs\<close>. | |
| 57349 | 2285 | |
| 2286 | Based on the dependency graph of tasks, the current thread takes over the | |
| 2287 | responsibility to evaluate future expressions that are required for the main | |
| 2288 | result, working from the bottom up. Waiting on future results that are | |
| 2289 | presently evaluated on other threads only happens as last resort, when no | |
| 2290 | other unfinished futures are left over. | |
| 2291 | ||
| 69597 | 2292 | \<^descr> \<^ML>\<open>Future.value\<close>~\<open>a\<close> wraps the value \<open>a\<close> as finished future value, | 
| 61854 | 2293 | bypassing the worker-thread farm. When joined, it returns \<open>a\<close> without any | 
| 2294 | further evaluation. | |
| 57349 | 2295 | |
| 2296 | There is very low overhead for this proforma wrapping of strict values as | |
| 57421 | 2297 | futures. | 
| 57348 | 2298 | |
| 75617 | 2299 | \<^descr> \<^ML>\<open>Future.map\<close>~\<open>f x\<close> is a fast-path implementation of | 
| 2300 | \<^ML>\<open>Future.fork\<close>~\<open>(fn () => f (\<close>\<^ML>\<open>Future.join\<close>~\<open>x))\<close>, which avoids | |
| 2301 | the full overhead of the task queue and worker-thread farm as far as | |
| 2302 | possible. The function \<open>f\<close> is supposed to be some trivial post-processing or | |
| 2303 | projection of the future result. | |
| 57348 | 2304 | |
| 69597 | 2305 | \<^descr> \<^ML>\<open>Future.cancel\<close>~\<open>x\<close> cancels the task group of the given future, using | 
| 2306 | \<^ML>\<open>Future.cancel_group\<close> below. | |
| 2307 | ||
| 2308 | \<^descr> \<^ML>\<open>Future.cancel_group\<close>~\<open>group\<close> cancels all tasks of the given task | |
| 61854 | 2309 | group for all time. Threads that are presently processing a task of the | 
| 2310 | given group are interrupted: it may take some time until they are actually | |
| 2311 | terminated. Tasks that are queued but not yet processed are dequeued and | |
| 2312 | forced into interrupted state. Since the task group is itself invalidated, | |
| 2313 | any further attempt to fork a future that belongs to it will yield a | |
| 2314 | canceled result as well. | |
| 57348 | 2315 | |
| 69597 | 2316 | \<^descr> \<^ML>\<open>Future.promise\<close>~\<open>abort\<close> registers a passive future with the given | 
| 61854 | 2317 | \<open>abort\<close> operation: it is invoked when the future task group is canceled. | 
| 57348 | 2318 | |
| 69597 | 2319 | \<^descr> \<^ML>\<open>Future.fulfill\<close>~\<open>x a\<close> finishes the passive future \<open>x\<close> by the given | 
| 61854 | 2320 | value \<open>a\<close>. If the promise has already been canceled, the attempt to fulfill | 
| 2321 | it causes an exception. | |
| 58618 | 2322 | \<close> | 
| 57348 | 2323 | |
| 47180 | 2324 | end |