author | wenzelm |
Mon, 21 Jul 2014 16:58:12 +0200 | |
changeset 57593 | 2f7d91242b99 |
parent 56618 | 874bdedb2313 |
child 58616 | 4257a7f2bf39 |
permissions | -rw-r--r-- |
56276 | 1 |
(* Title: Tools/SML/Examples.thy |
2 |
Author: Makarius |
|
3 |
*) |
|
4 |
||
5 |
header {* Standard ML within the Isabelle environment *} |
|
6 |
||
7 |
theory Examples |
|
8 |
imports Pure |
|
9 |
begin |
|
10 |
||
11 |
text {* |
|
12 |
Isabelle/ML is a somewhat augmented version of Standard ML, with |
|
13 |
various add-ons that are indispensable for Isabelle development, but |
|
14 |
may cause conflicts with independent projects in plain Standard ML. |
|
15 |
||
16 |
The Isabelle/Isar command 'SML_file' supports official Standard ML |
|
17 |
within the Isabelle environment, with full support in the Prover IDE |
|
18 |
(Isabelle/jEdit). |
|
19 |
||
20 |
Here is a very basic example that defines the factorial function and |
|
21 |
evaluates it for some arguments. |
|
22 |
*} |
|
23 |
||
24 |
SML_file "factorial.sml" |
|
25 |
||
26 |
text {* |
|
27 |
The subsequent example illustrates the use of multiple 'SML_file' |
|
28 |
commands to build larger Standard ML projects. The toplevel SML |
|
29 |
environment is enriched cumulatively within the theory context of |
|
30 |
Isabelle --- independently of the Isabelle/ML environment. |
|
31 |
*} |
|
32 |
||
33 |
SML_file "Example.sig" |
|
34 |
SML_file "Example.sml" |
|
35 |
||
56618
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
36 |
|
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
37 |
text {* |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
38 |
Isabelle/ML and SML share a common run-time system, but the static |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
39 |
environments are separate. It is possible to exchange toplevel bindings |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
40 |
via separate commands like this. |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
41 |
*} |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
42 |
|
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
43 |
SML_export {* val f = factorial *} -- {* re-use factorial from SML environment *} |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
44 |
ML {* f 42 *} |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
45 |
|
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
46 |
SML_import {* val println = Output.writeln *} |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
47 |
-- {* re-use Isabelle/ML channel for SML *} |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
48 |
|
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
49 |
SML_import {* val par_map = Par_List.map *} |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
50 |
-- {* re-use Isabelle/ML parallel list combinator for SML *} |
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents:
56276
diff
changeset
|
51 |
|
56276 | 52 |
end |