| author | wenzelm | 
| Mon, 18 Oct 2010 19:06:07 +0100 | |
| changeset 39866 | 5ec01d5acd0c | 
| parent 39864 | f3b4fde34cd1 | 
| child 39867 | a8363532cd4d | 
| permissions | -rw-r--r-- | 
| 29755 | 1  | 
theory "ML"  | 
2  | 
imports Base  | 
|
3  | 
begin  | 
|
| 18538 | 4  | 
|
| 39822 | 5  | 
chapter {* Isabelle/ML *}
 | 
| 20489 | 6  | 
|
| 39823 | 7  | 
text {* Isabelle/ML is best understood as a certain culture based on
 | 
8  | 
Standard ML. Thus it is not a new programming language, but a  | 
|
9  | 
certain way to use SML at an advanced level within the Isabelle  | 
|
10  | 
environment. This covers a variety of aspects that are geared  | 
|
11  | 
towards an efficient and robust platform for applications of formal  | 
|
12  | 
logic with fully foundational proof construction --- according to  | 
|
13  | 
  the well-known \emph{LCF principle}.  There are specific library
 | 
|
14  | 
modules and infrastructure to address the needs for such difficult  | 
|
15  | 
tasks. For example, the raw parallel programming model of Poly/ML  | 
|
16  | 
  is presented as considerably more abstract concept of \emph{future
 | 
|
17  | 
values}, which is then used to augment the inference kernel, proof  | 
|
18  | 
interpreter, and theory loader accordingly.  | 
|
19  | 
||
20  | 
The main aspects of Isabelle/ML are introduced below. These  | 
|
21  | 
first-hand explanations should help to understand how proper  | 
|
22  | 
Isabelle/ML is to be read and written, and to get access to the  | 
|
23  | 
wealth of experience that is expressed in the source text and its  | 
|
24  | 
  history of changes.\footnote{See
 | 
|
25  | 
  \url{http://isabelle.in.tum.de/repos/isabelle} for the full
 | 
|
26  | 
Mercurial history. There are symbolic tags to refer to official  | 
|
27  | 
  Isabelle releases, as opposed to arbitrary \emph{tip} versions that
 | 
|
28  | 
merely reflect snapshots that are never really up-to-date.} *}  | 
|
29  | 
||
30  | 
||
31  | 
section {* SML embedded into Isabelle/Isar *}
 | 
|
32  | 
||
| 39824 | 33  | 
text {* ML and Isar are intertwined via an open-ended bootstrap
 | 
34  | 
process that provides more and more programming facilities and  | 
|
35  | 
logical content in an alternating manner. Bootstrapping starts from  | 
|
36  | 
the raw environment of existing implementations of Standard ML  | 
|
37  | 
(mainly Poly/ML, but also SML/NJ).  | 
|
| 39823 | 38  | 
|
| 39824 | 39  | 
Isabelle/Pure marks the point where the original ML toplevel is  | 
40  | 
superseded by the Isar toplevel that maintains a uniform environment  | 
|
41  | 
  for arbitrary ML values (see also \secref{sec:context}).  This
 | 
|
42  | 
formal context holds logical entities as well as ML compiler  | 
|
43  | 
bindings, among many other things. Raw Standard ML is never  | 
|
44  | 
encountered again after the initial bootstrap of Isabelle/Pure.  | 
|
| 39823 | 45  | 
|
| 39824 | 46  | 
Object-logics such as Isabelle/HOL are built within the  | 
47  | 
Isabelle/ML/Isar environment of Pure by introducing suitable  | 
|
48  | 
theories with associated ML text, either inlined or as separate  | 
|
49  | 
files. Thus Isabelle/HOL is defined as a regular user-space  | 
|
50  | 
application within the Isabelle framework. Further add-on tools can  | 
|
51  | 
be implemented in ML within the Isar context in the same manner: ML  | 
|
52  | 
is part of the regular user-space repertoire of Isabelle.  | 
|
| 39823 | 53  | 
*}  | 
54  | 
||
| 39824 | 55  | 
|
| 
39827
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
56  | 
subsection {* Isar ML commands *}
 | 
| 39823 | 57  | 
|
58  | 
text {* The primary Isar source language provides various facilities
 | 
|
59  | 
to open a ``window'' to the underlying ML compiler. Especially see  | 
|
| 39824 | 60  | 
  @{command_ref "use"} and @{command_ref "ML"}, which work exactly the
 | 
61  | 
same way, only the source text is provided via a file vs.\ inlined,  | 
|
62  | 
respectively. Apart from embedding ML into the main theory  | 
|
63  | 
definition like that, there are many more commands that refer to ML  | 
|
64  | 
  source, such as @{command_ref setup} or @{command_ref declaration}.
 | 
|
65  | 
An example of even more fine-grained embedding of ML into Isar is  | 
|
66  | 
  the proof method @{method_ref tactic}, which refines the pending goal state
 | 
|
67  | 
  via a given expression of type @{ML_type tactic}.
 | 
|
68  | 
*}  | 
|
| 39823 | 69  | 
|
| 39824 | 70  | 
text %mlex {* The following artificial example demonstrates some ML
 | 
71  | 
toplevel declarations within the implicit Isar theory context. This  | 
|
72  | 
is regular functional programming without referring to logical  | 
|
73  | 
entities yet.  | 
|
| 39823 | 74  | 
*}  | 
75  | 
||
76  | 
ML {*
 | 
|
77  | 
fun factorial 0 = 1  | 
|
78  | 
| factorial n = n * factorial (n - 1)  | 
|
79  | 
*}  | 
|
80  | 
||
| 
39861
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
81  | 
text {* Here the ML environment is really managed by Isabelle, i.e.\
 | 
| 
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
82  | 
  the @{ML factorial} function is not yet accessible in the preceding
 | 
| 
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
83  | 
paragraph, nor in a different theory that is independent from the  | 
| 
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
84  | 
current one in the import hierarchy.  | 
| 39823 | 85  | 
|
86  | 
Removing the above ML declaration from the source text will remove  | 
|
87  | 
any trace of this definition as expected. The Isabelle/ML toplevel  | 
|
88  | 
  environment is managed in a \emph{stateless} way: unlike the raw ML
 | 
|
| 39824 | 89  | 
toplevel or similar command loops of Computer Algebra systems, there  | 
90  | 
  are no global side-effects involved here.\footnote{Such a stateless
 | 
|
91  | 
compilation environment is also a prerequisite for robust parallel  | 
|
92  | 
compilation within independent nodes of the implicit theory  | 
|
93  | 
development graph.}  | 
|
| 39823 | 94  | 
|
| 
39861
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
95  | 
\medskip The next example shows how to embed ML into Isar proofs.  | 
| 39824 | 96  | 
  Instead of @{command_ref "ML"} for theory mode, we use @{command_ref
 | 
97  | 
"ML_prf"} for proof mode. As illustrated below, its effect on the  | 
|
98  | 
ML environment is local to the whole proof body, while ignoring its  | 
|
| 
39861
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
99  | 
internal block structure. *}  | 
| 39823 | 100  | 
|
101  | 
example_proof  | 
|
| 39851 | 102  | 
  ML_prf %"ML" {* val a = 1 *}
 | 
| 39824 | 103  | 
  { -- {* Isar block structure ignored by ML environment *}
 | 
| 39851 | 104  | 
    ML_prf %"ML" {* val b = a + 1 *}
 | 
| 39824 | 105  | 
  } -- {* Isar block structure ignored by ML environment *}
 | 
| 39851 | 106  | 
  ML_prf %"ML" {* val c = b + 1 *}
 | 
| 39823 | 107  | 
qed  | 
108  | 
||
| 
39861
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
109  | 
text {* By side-stepping the normal scoping rules for Isar proof
 | 
| 
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
110  | 
blocks, embedded ML code can refer to the different contexts  | 
| 39824 | 111  | 
explicitly, and manipulate corresponding entities, e.g.\ export a  | 
112  | 
fact from a block context.  | 
|
| 39823 | 113  | 
|
| 
39861
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
114  | 
\medskip Two further ML commands are useful in certain situations:  | 
| 39824 | 115  | 
  @{command_ref ML_val} and @{command_ref ML_command} are both
 | 
116  | 
  \emph{diagnostic} in the sense that there is no effect on the
 | 
|
117  | 
underlying environment, and can thus used anywhere (even outside a  | 
|
118  | 
theory). The examples below produce long strings of digits by  | 
|
119  | 
  invoking @{ML factorial}: @{command ML_val} already takes care of
 | 
|
120  | 
  printing the ML toplevel result, but @{command ML_command} is silent
 | 
|
| 
39861
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
121  | 
so we produce an explicit output message. *}  | 
| 39823 | 122  | 
|
123  | 
ML_val {* factorial 100 *}
 | 
|
124  | 
ML_command {* writeln (string_of_int (factorial 100)) *}
 | 
|
125  | 
||
126  | 
example_proof  | 
|
| 39824 | 127  | 
  ML_val {* factorial 100 *}  (* FIXME check/fix indentation *)
 | 
| 39823 | 128  | 
  ML_command {* writeln (string_of_int (factorial 100)) *}
 | 
129  | 
qed  | 
|
130  | 
||
131  | 
||
| 
39827
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
132  | 
subsection {* Compile-time context *}
 | 
| 39823 | 133  | 
|
| 
39825
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
134  | 
text {* Whenever the ML compiler is invoked within Isabelle/Isar, the
 | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
135  | 
formal context is passed as a thread-local reference variable. Thus  | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
136  | 
ML code may access the theory context during compilation, by reading  | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
137  | 
or writing the (local) theory under construction. Note that such  | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
138  | 
direct access to the compile-time context is rare; in practice it is  | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
139  | 
typically via some derived ML functions.  | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
140  | 
*}  | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
141  | 
|
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
142  | 
text %mlref {*
 | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
143  | 
  \begin{mldecls}
 | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
144  | 
  @{index_ML ML_Context.the_generic_context: "unit -> Context.generic"} \\
 | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
145  | 
  @{index_ML "Context.>> ": "(Context.generic -> Context.generic) -> unit"} \\
 | 
| 39850 | 146  | 
  @{index_ML bind_thms: "string * thm list -> unit"} \\
 | 
147  | 
  @{index_ML bind_thm: "string * thm -> unit"} \\
 | 
|
| 
39825
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
148  | 
  \end{mldecls}
 | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
149  | 
|
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
150  | 
  \begin{description}
 | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
151  | 
|
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
152  | 
  \item @{ML "ML_Context.the_generic_context ()"} refers to the theory
 | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
153  | 
context of the ML toplevel --- at compile time. ML code needs to  | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
154  | 
  take care to refer to @{ML "ML_Context.the_generic_context ()"}
 | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
155  | 
correctly. Recall that evaluation of a function body is delayed  | 
| 
39827
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
156  | 
until actual run-time.  | 
| 
39825
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
157  | 
|
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
158  | 
  \item @{ML "Context.>>"}~@{text f} applies context transformation
 | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
159  | 
  @{text f} to the implicit context of the ML toplevel.
 | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
160  | 
|
| 39850 | 161  | 
  \item @{ML bind_thms}~@{text "(name, thms)"} stores a list of
 | 
162  | 
theorems produced in ML both in the (global) theory context and the  | 
|
163  | 
ML toplevel, associating it with the provided name. Theorems are  | 
|
164  | 
put into a global ``standard'' format before being stored.  | 
|
165  | 
||
166  | 
  \item @{ML bind_thm} is similar to @{ML bind_thms} but refers to a
 | 
|
167  | 
singleton theorem.  | 
|
168  | 
||
| 
39825
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
169  | 
  \end{description}
 | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
170  | 
|
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
171  | 
It is very important to note that the above functions are really  | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
172  | 
restricted to the compile time, even though the ML compiler is  | 
| 
39827
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
173  | 
invoked at run-time. The majority of ML code either uses static  | 
| 
39825
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
174  | 
  antiquotations (\secref{sec:ML-antiq}) or refers to the theory or
 | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
175  | 
proof context at run-time, by explicit functional abstraction.  | 
| 
 
f9066b94bf07
eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
 
wenzelm 
parents: 
39824 
diff
changeset
 | 
176  | 
*}  | 
| 39823 | 177  | 
|
178  | 
||
| 
39827
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
179  | 
subsection {* Antiquotations \label{sec:ML-antiq} *}
 | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
180  | 
|
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
181  | 
text {* A very important consequence of embedding SML into Isar is the
 | 
| 39829 | 182  | 
  concept of \emph{ML antiquotation}.  First, the standard token
 | 
183  | 
language of ML is augmented by special syntactic entities of the  | 
|
184  | 
following form:  | 
|
| 
39827
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
185  | 
|
| 39829 | 186  | 
  \indexouternonterm{antiquote}
 | 
| 
39827
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
187  | 
  \begin{rail}
 | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
188  | 
antiquote: atsign lbrace nameref args rbrace | lbracesym | rbracesym  | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
189  | 
;  | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
190  | 
  \end{rail}
 | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
191  | 
|
| 
39861
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
192  | 
  Here @{syntax nameref} and @{syntax args} are regular outer syntax
 | 
| 
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
193  | 
categories. Note that attributes and proof methods use similar  | 
| 
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
194  | 
syntax.  | 
| 39823 | 195  | 
|
| 
39827
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
196  | 
  \medskip A regular antiquotation @{text "@{name args}"} processes
 | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
197  | 
its arguments by the usual means of the Isar source language, and  | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
198  | 
produces corresponding ML source text, either as literal  | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
199  | 
  \emph{inline} text (e.g. @{text "@{term t}"}) or abstract
 | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
200  | 
  \emph{value} (e.g. @{text "@{thm th}"}).  This pre-compilation
 | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
201  | 
scheme allows to refer to formal entities in a robust manner, with  | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
202  | 
proper static scoping and with some degree of logical checking of  | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
203  | 
small portions of the code.  | 
| 39823 | 204  | 
|
| 
39827
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
205  | 
  Special antiquotations like @{text "@{let \<dots>}"} or @{text "@{note
 | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
206  | 
\<dots>}"} augment the compilation context without generating code. The  | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
207  | 
  non-ASCII braces @{text "\<lbrace>"} and @{text "\<rbrace>"} allow to delimit the
 | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
208  | 
effect by introducing local blocks within the pre-compilation  | 
| 
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
209  | 
environment.  | 
| 39829 | 210  | 
|
| 
39861
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
211  | 
  \medskip See also \cite{Wenzel-Chaieb:2007b} for a slightly broader
 | 
| 
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
212  | 
perspective on Isabelle/ML antiquotations. *}  | 
| 39829 | 213  | 
|
| 39830 | 214  | 
text %mlantiq {*
 | 
| 39829 | 215  | 
  \begin{matharray}{rcl}
 | 
216  | 
  @{ML_antiquotation_def "let"} & : & @{text ML_antiquotation} \\
 | 
|
217  | 
  @{ML_antiquotation_def "note"} & : & @{text ML_antiquotation} \\
 | 
|
218  | 
  \end{matharray}
 | 
|
219  | 
||
220  | 
  \begin{rail}
 | 
|
221  | 
'let' ((term + 'and') '=' term + 'and')  | 
|
222  | 
;  | 
|
223  | 
||
224  | 
'note' (thmdef? thmrefs + 'and')  | 
|
225  | 
;  | 
|
226  | 
  \end{rail}
 | 
|
227  | 
||
228  | 
  \begin{description}
 | 
|
229  | 
||
230  | 
  \item @{text "@{let p = t}"} binds schematic variables in the
 | 
|
231  | 
  pattern @{text "p"} by higher-order matching against the term @{text
 | 
|
232  | 
  "t"}.  This is analogous to the regular @{command_ref let} command
 | 
|
233  | 
in the Isar proof language. The pre-compilation environment is  | 
|
234  | 
augmented by auxiliary term bindings, without emitting ML source.  | 
|
235  | 
||
236  | 
  \item @{text "@{note a = b\<^sub>1 \<dots> b\<^sub>n}"} recalls existing facts @{text
 | 
|
237  | 
  "b\<^sub>1, \<dots>, b\<^sub>n"}, binding the result as @{text a}.  This is analogous to
 | 
|
238  | 
  the regular @{command_ref note} command in the Isar proof language.
 | 
|
239  | 
The pre-compilation environment is augmented by auxiliary fact  | 
|
240  | 
bindings, without emitting ML source.  | 
|
241  | 
||
242  | 
  \end{description}
 | 
|
243  | 
*}  | 
|
244  | 
||
245  | 
text %mlex {* The following artificial example gives a first
 | 
|
246  | 
impression about using the antiquotation elements introduced so far,  | 
|
247  | 
  together with the basic @{text "@{thm}"} antiquotation defined
 | 
|
248  | 
later.  | 
|
249  | 
*}  | 
|
250  | 
||
251  | 
ML {*
 | 
|
252  | 
\<lbrace>  | 
|
253  | 
    @{let ?t = my_term}
 | 
|
254  | 
    @{note my_refl = reflexive [of ?t]}
 | 
|
255  | 
    fun foo th = Thm.transitive th @{thm my_refl}
 | 
|
256  | 
\<rbrace>  | 
|
257  | 
*}  | 
|
258  | 
||
259  | 
text {*
 | 
|
260  | 
The extra block delimiters do not affect the compiled code itself, i.e.\  | 
|
261  | 
  function @{ML foo} is available in the present context.
 | 
|
| 
39827
 
d829ce302ca4
basic setup for ML antiquotations -- with rail diagrams;
 
wenzelm 
parents: 
39825 
diff
changeset
 | 
262  | 
*}  | 
| 39823 | 263  | 
|
| 39835 | 264  | 
|
| 39854 | 265  | 
section {* Message output channels \label{sec:message-channels} *}
 | 
| 39835 | 266  | 
|
267  | 
text {* Isabelle provides output channels for different kinds of
 | 
|
268  | 
messages: regular output, high-volume tracing information, warnings,  | 
|
269  | 
and errors.  | 
|
270  | 
||
271  | 
Depending on the user interface involved, these messages may appear  | 
|
272  | 
in different text styles or colours. The standard output for  | 
|
273  | 
  terminal sessions prefixes each line of warnings by @{verbatim
 | 
|
274  | 
  "###"} and errors by @{verbatim "***"}, but leaves anything else
 | 
|
275  | 
unchanged.  | 
|
276  | 
||
277  | 
Messages are associated with the transaction context of the running  | 
|
278  | 
Isar command. This enables the front-end to manage commands and  | 
|
279  | 
resulting messages together. For example, after deleting a command  | 
|
280  | 
from a given theory document version, the corresponding message  | 
|
281  | 
output can be retracted from the display. *}  | 
|
282  | 
||
283  | 
text %mlref {*
 | 
|
284  | 
  \begin{mldecls}
 | 
|
285  | 
  @{index_ML writeln: "string -> unit"} \\
 | 
|
286  | 
  @{index_ML tracing: "string -> unit"} \\
 | 
|
287  | 
  @{index_ML warning: "string -> unit"} \\
 | 
|
288  | 
  @{index_ML error: "string -> 'a"} \\
 | 
|
289  | 
  \end{mldecls}
 | 
|
290  | 
||
291  | 
  \begin{description}
 | 
|
292  | 
||
293  | 
  \item @{ML writeln}~@{text "text"} outputs @{text "text"} as regular
 | 
|
294  | 
message. This is the primary message output operation of Isabelle  | 
|
295  | 
and should be used by default.  | 
|
296  | 
||
297  | 
  \item @{ML tracing}~@{text "text"} outputs @{text "text"} as special
 | 
|
298  | 
tracing message, indicating potential high-volume output to the  | 
|
299  | 
front-end (hundreds or thousands of messages issued by a single  | 
|
300  | 
command). The idea is to allow the user-interface to downgrade the  | 
|
301  | 
quality of message display to achieve higher throughput.  | 
|
302  | 
||
303  | 
Note that the user might have to take special actions to see tracing  | 
|
304  | 
output, e.g.\ switch to a different output window. So this channel  | 
|
305  | 
should not be used for regular output.  | 
|
306  | 
||
307  | 
  \item @{ML warning}~@{text "text"} outputs @{text "text"} as
 | 
|
308  | 
warning, which typically means some extra emphasis on the front-end  | 
|
309  | 
side (color highlighting, icon).  | 
|
310  | 
||
311  | 
  \item @{ML error}~@{text "text"} raises exception @{ML ERROR}~@{text
 | 
|
312  | 
  "text"} and thus lets the Isar toplevel print @{text "text"} on the
 | 
|
313  | 
error channel, which typically means some extra emphasis on the  | 
|
314  | 
front-end side (color highlighting, icon).  | 
|
315  | 
||
316  | 
This assumes that the exception is not handled before the command  | 
|
317  | 
  terminates.  Handling exception @{ML ERROR}~@{text "text"} is a
 | 
|
318  | 
perfectly legal alternative: it means that the error is absorbed  | 
|
319  | 
without any message output.  | 
|
320  | 
||
| 
39861
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
321  | 
  \begin{warn}
 | 
| 39835 | 322  | 
  The actual error channel is accessed via @{ML Output.error_msg}, but
 | 
323  | 
  the interaction protocol of Proof~General \emph{crashes} if that
 | 
|
324  | 
function is used in regular ML code: error output and toplevel  | 
|
325  | 
command failure always need to coincide here.  | 
|
| 
39861
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
326  | 
  \end{warn}
 | 
| 39835 | 327  | 
|
| 
39861
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
328  | 
  \end{description}
 | 
| 
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
329  | 
|
| 
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
330  | 
  \begin{warn}
 | 
| 39835 | 331  | 
Regular Isabelle/ML code should output messages exclusively by the  | 
332  | 
  official channels.  Using raw I/O on \emph{stdout} or \emph{stderr}
 | 
|
333  | 
  instead (e.g.\ via @{ML TextIO.output}) is apt to cause problems in
 | 
|
334  | 
the presence of parallel and asynchronous processing of Isabelle  | 
|
335  | 
theories. Such raw output might be displayed by the front-end in  | 
|
336  | 
some system console log, with a low chance that the user will ever  | 
|
337  | 
see it. Moreover, as a genuine side-effect on global process  | 
|
338  | 
channels, there is no proper way to retract output when Isar command  | 
|
339  | 
transactions are reset.  | 
|
| 
39861
 
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
 
wenzelm 
parents: 
39859 
diff
changeset
 | 
340  | 
  \end{warn}
 | 
| 39835 | 341  | 
*}  | 
342  | 
||
| 39854 | 343  | 
|
344  | 
section {* Exceptions *}
 | 
|
345  | 
||
346  | 
text {* The Standard ML semantics of strict functional evaluation
 | 
|
347  | 
together with exceptions is rather well defined, but some delicate  | 
|
348  | 
points need to be observed to avoid that ML programs go wrong  | 
|
349  | 
despite static type-checking. Exceptions in Isabelle/ML are  | 
|
350  | 
subsequently categorized as follows.  | 
|
351  | 
||
352  | 
  \paragraph{Regular user errors.}  These are meant to provide
 | 
|
353  | 
informative feedback about malformed input etc.  | 
|
354  | 
||
355  | 
  The \emph{error} function raises the corresponding \emph{ERROR}
 | 
|
356  | 
  exception, with a plain text message as argument.  \emph{ERROR}
 | 
|
357  | 
exceptions can be handled internally, in order to be ignored, turned  | 
|
358  | 
into other exceptions, or cascaded by appending messages. If the  | 
|
359  | 
  corresponding Isabelle/Isar command terminates with an \emph{ERROR}
 | 
|
| 39855 | 360  | 
exception state, the toplevel will print the result on the error  | 
361  | 
  channel (see \secref{sec:message-channels}).
 | 
|
| 39854 | 362  | 
|
363  | 
It is considered bad style to refer to internal function names or  | 
|
364  | 
values in ML source notation in user error messages.  | 
|
365  | 
||
366  | 
Grammatical correctness of error messages can be improved by  | 
|
367  | 
  \emph{omitting} final punctuation: messages are often concatenated
 | 
|
368  | 
or put into a larger context (e.g.\ augmented with source position).  | 
|
369  | 
By not insisting in the final word at the origin of the error, the  | 
|
370  | 
system can perform its administrative tasks more easily and  | 
|
371  | 
robustly.  | 
|
372  | 
||
373  | 
  \paragraph{Program failures.}  There is a handful of standard
 | 
|
374  | 
exceptions that indicate general failure situations, or failures of  | 
|
375  | 
core operations on logical entities (types, terms, theorems,  | 
|
| 39856 | 376  | 
  theories, see \chref{ch:logic}).
 | 
| 39854 | 377  | 
|
378  | 
These exceptions indicate a genuine breakdown of the program, so the  | 
|
379  | 
main purpose is to determine quickly what has happened where.  | 
|
| 39855 | 380  | 
Traditionally, the (short) exception message would include the name  | 
381  | 
of an ML function, although this no longer necessary, because the ML  | 
|
382  | 
runtime system prints a detailed source position of the  | 
|
383  | 
  corresponding @{verbatim raise} keyword.
 | 
|
| 39854 | 384  | 
|
385  | 
\medskip User modules can always introduce their own custom  | 
|
386  | 
exceptions locally, e.g.\ to organize internal failures robustly  | 
|
387  | 
without overlapping with existing exceptions. Exceptions that are  | 
|
388  | 
exposed in module signatures require extra care, though, and should  | 
|
389  | 
  \emph{not} be introduced by default.  Surprise by end-users or ML
 | 
|
| 39855 | 390  | 
users of a module can be often minimized by using plain user errors.  | 
| 39854 | 391  | 
|
392  | 
  \paragraph{Interrupts.}  These indicate arbitrary system events:
 | 
|
393  | 
both the ML runtime system and the Isabelle/ML infrastructure signal  | 
|
394  | 
various exceptional situations by raising the special  | 
|
395  | 
  \emph{Interrupt} exception in user code.
 | 
|
396  | 
||
397  | 
This is the one and only way that physical events can intrude an  | 
|
398  | 
Isabelle/ML program. Such an interrupt can mean out-of-memory,  | 
|
399  | 
stack overflow, timeout, internal signaling of threads, or the user  | 
|
| 39855 | 400  | 
producing a console interrupt manually etc. An Isabelle/ML program  | 
401  | 
that intercepts interrupts becomes dependent on physical effects of  | 
|
402  | 
the environment. Even worse, exception handling patterns that are  | 
|
403  | 
too general by accident, e.g.\ by mispelled exception constructors,  | 
|
404  | 
will cover interrupts unintentionally, and thus render the program  | 
|
405  | 
semantics ill-defined.  | 
|
| 39854 | 406  | 
|
407  | 
Note that the Interrupt exception dates back to the original SML90  | 
|
408  | 
language definition. It was excluded from the SML97 version to  | 
|
409  | 
avoid its malign impact on ML program semantics, but without  | 
|
410  | 
providing a viable alternative. Isabelle/ML recovers physical  | 
|
411  | 
interruptibility (which an indispensable tool to implement managed  | 
|
412  | 
evaluation of Isar command transactions), but requires user code to  | 
|
413  | 
be strictly transparent wrt.\ interrupts.  | 
|
414  | 
||
415  | 
  \begin{warn}
 | 
|
416  | 
Isabelle/ML user code needs to terminate promptly on interruption,  | 
|
417  | 
without guessing at its meaning to the system infrastructure.  | 
|
418  | 
Temporary handling of interrupts for cleanup of global resources  | 
|
419  | 
etc.\ needs to be followed immediately by re-raising of the original  | 
|
420  | 
exception.  | 
|
421  | 
  \end{warn}
 | 
|
422  | 
*}  | 
|
423  | 
||
| 39855 | 424  | 
text %mlref {*
 | 
425  | 
  \begin{mldecls}
 | 
|
426  | 
  @{index_ML try: "('a -> 'b) -> 'a -> 'b option"} \\
 | 
|
427  | 
  @{index_ML can: "('a -> 'b) -> 'a -> bool"} \\
 | 
|
| 39856 | 428  | 
  @{index_ML ERROR: "string -> exn"} \\
 | 
429  | 
  @{index_ML Fail: "string -> exn"} \\
 | 
|
430  | 
  @{index_ML Exn.is_interrupt: "exn -> bool"} \\
 | 
|
| 39855 | 431  | 
  @{index_ML reraise: "exn -> 'a"} \\
 | 
432  | 
  @{index_ML exception_trace: "(unit -> 'a) -> 'a"} \\
 | 
|
433  | 
  \end{mldecls}
 | 
|
434  | 
||
435  | 
  \begin{description}
 | 
|
436  | 
||
437  | 
  \item @{ML try}~@{text "f x"} makes the partiality of evaluating
 | 
|
438  | 
  @{text "f x"} explicit via the option datatype.  Interrupts are
 | 
|
439  | 
  \emph{not} handled here, i.e.\ this form serves as safe replacement
 | 
|
440  | 
  for the \emph{unsafe} version @{verbatim "(SOME"}~@{text "f
 | 
|
441  | 
  x"}~@{verbatim "handle _ => NONE)"} that is occasionally seen in
 | 
|
442  | 
books about SML.  | 
|
443  | 
||
444  | 
  \item @{ML can} is similar to @{ML try} with more abstract result.
 | 
|
445  | 
||
| 39856 | 446  | 
  \item @{ML ERROR}~@{text "msg"} represents user errors; this
 | 
447  | 
  exception is always raised via the @{ML error} function (see
 | 
|
448  | 
  \secref{sec:message-channels}).
 | 
|
449  | 
||
450  | 
  \item @{ML Fail}~@{text "msg"} represents general program failures.
 | 
|
451  | 
||
452  | 
  \item @{ML Exn.is_interrupt} identifies interrupts robustly, without
 | 
|
453  | 
mentioning concrete exception constructors in user code. Handled  | 
|
454  | 
interrupts need to be re-raised promptly!  | 
|
455  | 
||
| 39855 | 456  | 
  \item @{ML reraise}~@{text "exn"} raises exception @{text "exn"}
 | 
457  | 
while preserving its implicit position information (if possible,  | 
|
458  | 
depending on the ML platform).  | 
|
459  | 
||
460  | 
  \item @{ML exception_trace}~@{verbatim "(fn _ =>"}~@{text
 | 
|
461  | 
  "e"}@{verbatim ")"} evaluates expression @{text "e"} while printing
 | 
|
462  | 
a full trace of its stack of nested exceptions (if possible,  | 
|
463  | 
  depending on the ML platform).\footnote{In various versions of
 | 
|
464  | 
Poly/ML the trace will appear on raw stdout of the Isabelle  | 
|
465  | 
process.}  | 
|
466  | 
||
467  | 
  Inserting @{ML exception_trace} into ML code temporarily is useful
 | 
|
468  | 
for debugging, but not suitable for production code.  | 
|
469  | 
||
470  | 
  \end{description}
 | 
|
471  | 
*}  | 
|
472  | 
||
| 
39866
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
473  | 
text %mlantiq {*
 | 
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
474  | 
  \begin{matharray}{rcl}
 | 
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
475  | 
  @{ML_antiquotation_def "assert"} & : & @{text ML_antiquotation} \\
 | 
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
476  | 
  \end{matharray}
 | 
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
477  | 
|
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
478  | 
  \begin{description}
 | 
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
479  | 
|
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
480  | 
  \item @{text "@{assert}"} inlines a function @{text "bool \<Rightarrow> unit"}
 | 
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
481  | 
  that raises @{ML Fail} if the argument is @{ML false}.  Due to
 | 
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
482  | 
inlining the source position of failed assertions is included in the  | 
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
483  | 
error output.  | 
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
484  | 
|
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
485  | 
  \end{description}
 | 
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
486  | 
*}  | 
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
487  | 
|
| 39859 | 488  | 
|
| 39863 | 489  | 
section {* Basic data types *}
 | 
| 39859 | 490  | 
|
491  | 
text {* The basis library proposal of SML97 need to be treated with
 | 
|
492  | 
caution. Many of its operations simply do not fit with important  | 
|
493  | 
Isabelle/ML conventions (like ``canonical argument order'', see  | 
|
494  | 
  \secref{sec:canonical-argument-order}), others can cause serious
 | 
|
495  | 
problems with the parallel evaluation model of Isabelle/ML (such as  | 
|
496  | 
  @{ML TextIO.print} or @{ML OS.Process.system}).
 | 
|
497  | 
||
498  | 
Subsequently we give a brief overview of important operations on  | 
|
499  | 
basic ML data types.  | 
|
500  | 
*}  | 
|
501  | 
||
502  | 
||
| 39863 | 503  | 
subsection {* Characters *}
 | 
504  | 
||
505  | 
text %mlref {*
 | 
|
506  | 
  \begin{mldecls}
 | 
|
507  | 
  @{index_ML_type char} \\
 | 
|
508  | 
  \end{mldecls}
 | 
|
509  | 
||
510  | 
  \begin{description}
 | 
|
511  | 
||
| 39864 | 512  | 
  \item Type @{ML_type char} is \emph{not} used.  The smallest textual
 | 
513  | 
unit in Isabelle is represented a ``symbol'' (see  | 
|
514  | 
  \secref{sec:symbols}).
 | 
|
| 39863 | 515  | 
|
516  | 
  \end{description}
 | 
|
517  | 
*}  | 
|
518  | 
||
519  | 
||
| 39862 | 520  | 
subsection {* Integers *}
 | 
521  | 
||
522  | 
text %mlref {*
 | 
|
523  | 
  \begin{mldecls}
 | 
|
524  | 
  @{index_ML_type int} \\
 | 
|
525  | 
  \end{mldecls}
 | 
|
526  | 
||
527  | 
  \begin{description}
 | 
|
528  | 
||
| 39864 | 529  | 
  \item Type @{ML_type int} represents regular mathematical integers,
 | 
530  | 
  which are \emph{unbounded}.  Overflow never happens in
 | 
|
| 39862 | 531  | 
  practice.\footnote{The size limit for integer bit patterns in memory
 | 
532  | 
is 64\,MB for 32-bit Poly/ML, and much higher for 64-bit systems.}  | 
|
533  | 
This works uniformly for all supported ML platforms (Poly/ML and  | 
|
534  | 
SML/NJ).  | 
|
535  | 
||
536  | 
  Literal integers in ML text (e.g.\ @{ML
 | 
|
537  | 
123456789012345678901234567890}) are forced to be of this one true  | 
|
538  | 
integer type --- overloading of SML97 is disabled.  | 
|
539  | 
||
540  | 
  Structure @{ML_struct IntInf} of SML97 is obsolete, always use
 | 
|
541  | 
  @{ML_struct Int}.  Structure @{ML_struct Integer} in @{"file"
 | 
|
542  | 
"~~/src/Pure/General/integer.ML"} provides some additional  | 
|
543  | 
operations.  | 
|
544  | 
||
545  | 
  \end{description}
 | 
|
546  | 
*}  | 
|
547  | 
||
548  | 
||
| 39859 | 549  | 
subsection {* Options *}
 | 
550  | 
||
551  | 
text %mlref {*
 | 
|
552  | 
  \begin{mldecls}
 | 
|
553  | 
  @{index_ML Option.map: "('a -> 'b) -> 'a option -> 'b option"} \\
 | 
|
554  | 
  @{index_ML is_some: "'a option -> bool"} \\
 | 
|
555  | 
  @{index_ML is_none: "'a option -> bool"} \\
 | 
|
556  | 
  @{index_ML the: "'a option -> 'a"} \\
 | 
|
557  | 
  @{index_ML these: "'a list option -> 'a list"} \\
 | 
|
558  | 
  @{index_ML the_list: "'a option -> 'a list"} \\
 | 
|
559  | 
  @{index_ML the_default: "'a -> 'a option -> 'a"} \\
 | 
|
560  | 
  \end{mldecls}
 | 
|
561  | 
*}  | 
|
562  | 
||
563  | 
text {* Apart from @{ML Option.map} most operations defined in
 | 
|
564  | 
  structure @{ML_struct Option} are alien to Isabelle/ML.  The
 | 
|
565  | 
  operations shown above are defined in @{"file"
 | 
|
566  | 
"~~/src/Pure/General/basics.ML"}, among others. *}  | 
|
567  | 
||
568  | 
||
| 39863 | 569  | 
subsection {* Lists *}
 | 
570  | 
||
571  | 
text {* Lists are ubiquitous in ML as simple and light-weight
 | 
|
572  | 
``collections'' for many everyday programming tasks. Isabelle/ML  | 
|
573  | 
provides some important refinements over the predefined operations  | 
|
574  | 
in SML97. *}  | 
|
575  | 
||
576  | 
text %mlref {*
 | 
|
577  | 
  \begin{mldecls}
 | 
|
578  | 
  @{index_ML cons: "'a -> 'a list -> 'a list"} \\
 | 
|
579  | 
  \end{mldecls}
 | 
|
580  | 
||
581  | 
  \begin{description}
 | 
|
582  | 
||
583  | 
  \item @{ML cons}~@{text "x xs"} evaluates to @{text "x :: xs"}.
 | 
|
584  | 
||
585  | 
Tupled infix operators are a historical accident in Standard ML.  | 
|
586  | 
  The curried @{ML cons} amends this, but it should be only used when
 | 
|
587  | 
partial application is required.  | 
|
588  | 
||
589  | 
  \end{description}
 | 
|
590  | 
*}  | 
|
591  | 
||
592  | 
||
593  | 
subsubsection {* Canonical iteration *}
 | 
|
594  | 
||
595  | 
text {* A function @{text "f: \<alpha> \<rightarrow> \<beta> \<rightarrow> \<beta>"} can be understood as update
 | 
|
596  | 
  on a configuration of type @{text "\<beta>"} that is parametrized by
 | 
|
597  | 
  arguments of type @{text "\<alpha>"}.  Given @{text "a: \<alpha>"} the partial
 | 
|
598  | 
  application @{text "(f a): \<beta> \<rightarrow> \<beta>"} operates homogeneously on @{text
 | 
|
599  | 
"\<beta>"}. This can be iterated naturally over a list of parameters  | 
|
600  | 
  @{text "[a\<^sub>1, \<dots>, a\<^sub>n]"} as @{text "f a\<^sub>1; \<dots>; f a\<^bsub>n\<^esub>\<^bsub>\<^esub>"} (where the
 | 
|
601  | 
semicolon represents left-to-right composition). The latter  | 
|
602  | 
  expression is again a function @{text "\<beta> \<rightarrow> \<beta>"}.  It can be applied
 | 
|
603  | 
  to an initial configuration @{text "b: \<beta>"} to start the iteration
 | 
|
604  | 
  over the given list of arguments: each @{text "a"} in @{text "a\<^sub>1, \<dots>,
 | 
|
605  | 
a\<^sub>n"} is applied consecutively by updating a cumulative  | 
|
606  | 
configuration.  | 
|
607  | 
||
608  | 
  The @{text fold} combinator in Isabelle/ML lifts a function @{text
 | 
|
609  | 
"f"} as above to its iterated version over a list of arguments.  | 
|
610  | 
  Lifting can be repeated, e.g.\ @{text "(fold \<circ> fold) f"} iterates
 | 
|
611  | 
over a list of lists as expected.  | 
|
612  | 
||
613  | 
  The variant @{text "fold_rev"} works inside-out over the list of
 | 
|
614  | 
  arguments, such that @{text "fold_rev f \<equiv> fold f \<circ> rev"} holds.
 | 
|
615  | 
||
616  | 
  The @{text "fold_map"} combinator essentially performs @{text
 | 
|
617  | 
  "fold"} and @{text "map"} simultaneously: each application of @{text
 | 
|
618  | 
"f"} produces an updated configuration together with a side-result;  | 
|
619  | 
the iteration collects all such side-results as a separate list.  | 
|
620  | 
*}  | 
|
621  | 
||
622  | 
text %mlref {*
 | 
|
623  | 
  \begin{mldecls}
 | 
|
624  | 
  @{index_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
 | 
|
625  | 
  @{index_ML fold_rev: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
 | 
|
626  | 
  @{index_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\
 | 
|
627  | 
  \end{mldecls}
 | 
|
628  | 
||
629  | 
  \begin{description}
 | 
|
630  | 
||
631  | 
  \item @{ML fold}~@{text f} lifts the parametrized update function
 | 
|
632  | 
  @{text "f"} to a list of parameters.
 | 
|
633  | 
||
634  | 
  \item @{ML fold_rev}~@{text "f"} is similar to @{ML fold}~@{text
 | 
|
635  | 
"f"}, but works inside-out.  | 
|
636  | 
||
637  | 
  \item @{ML fold_map}~@{text "f"} lifts the parametrized update
 | 
|
638  | 
  function @{text "f"} (with side-result) to a list of parameters and
 | 
|
639  | 
cumulative side-results.  | 
|
640  | 
||
641  | 
  \end{description}
 | 
|
642  | 
||
643  | 
  \begin{warn}
 | 
|
644  | 
The literature on functional programming provides a multitude of  | 
|
645  | 
  combinators called @{text "foldl"}, @{text "foldr"} etc.  SML97
 | 
|
646  | 
  provides its own variations as @{ML List.foldl} and @{ML
 | 
|
647  | 
List.foldr}, while the classic Isabelle library still has the  | 
|
648  | 
  slightly more convenient historic @{ML Library.foldl} and @{ML
 | 
|
649  | 
Library.foldr}. To avoid further confusion, all of this should be  | 
|
650  | 
  ignored, and @{ML fold} (or @{ML fold_rev}) used exclusively.
 | 
|
651  | 
  \end{warn}
 | 
|
652  | 
*}  | 
|
653  | 
||
654  | 
text %mlex {* Using canonical @{ML fold} together with canonical @{ML
 | 
|
655  | 
cons}, or similar standard operations, alternates the orientation of  | 
|
656  | 
data. The is quite natural and should not altered forcible by  | 
|
657  | 
  inserting extra applications @{ML rev}.  The alternative @{ML
 | 
|
658  | 
fold_rev} can be used in the relatively few situations, where  | 
|
659  | 
alternation should be prevented.  | 
|
660  | 
*}  | 
|
661  | 
||
662  | 
ML {*
 | 
|
663  | 
val items = [1, 2, 3, 4, 5, 6, 7, 8, 9, 10];  | 
|
664  | 
||
665  | 
val list1 = fold cons items [];  | 
|
| 
39866
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
666  | 
  @{assert} (list1 = rev items);
 | 
| 
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
667  | 
|
| 39863 | 668  | 
val list2 = fold_rev cons items [];  | 
| 
39866
 
5ec01d5acd0c
more robust examples: explicit @{assert} instead of unchecked output;
 
wenzelm 
parents: 
39864 
diff
changeset
 | 
669  | 
  @{assert} (list2 = items);
 | 
| 39863 | 670  | 
*}  | 
671  | 
||
672  | 
||
| 39859 | 673  | 
subsection {* Unsynchronized references *}
 | 
674  | 
||
675  | 
text %mlref {*
 | 
|
676  | 
  \begin{mldecls}
 | 
|
677  | 
  @{index_ML Unsynchronized.ref: "'a -> 'a Unsynchronized.ref"} \\
 | 
|
678  | 
  @{index_ML "!": "'a Unsynchronized.ref -> 'a"} \\
 | 
|
679  | 
  @{index_ML "op :=": "'a Unsynchronized.ref * 'a -> unit"} \\
 | 
|
680  | 
  \end{mldecls}
 | 
|
681  | 
*}  | 
|
682  | 
||
683  | 
text {* Due to ubiquitous parallelism in Isabelle/ML (see also
 | 
|
684  | 
  \secref{sec:multi-threading}), the mutable reference cells of
 | 
|
685  | 
Standard ML are notorious for causing problems. In a highly  | 
|
686  | 
  parallel system, both correctness \emph{and} performance are easily
 | 
|
687  | 
degraded when using mutable data.  | 
|
688  | 
||
689  | 
  The unwieldy name of @{ML Unsynchronized.ref} for the constructor
 | 
|
690  | 
for references in Isabelle/ML emphasizes the inconveniences caused by  | 
|
691  | 
  mutability.  Existing operations @{ML "!"}  and @{ML "op :="} are
 | 
|
692  | 
unchanged, but should be used with special precautions, say in a  | 
|
693  | 
strictly local situation that is guaranteed to be restricted to  | 
|
694  | 
sequential evaluation -- now and in the future. *}  | 
|
695  | 
||
| 
30270
 
61811c9224a6
dummy changes to produce a new changeset of these files;
 
wenzelm 
parents: 
29755 
diff
changeset
 | 
696  | 
end  |