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