| author | wenzelm | 
| Sun, 06 Jan 2019 15:04:34 +0100 | |
| changeset 69605 | a96320074298 | 
| parent 61757 | 0d399131008f | 
| child 71934 | 914baafb3da4 | 
| permissions | -rw-r--r-- | 
| 56276 | 1  | 
(* Title: Tools/SML/Examples.thy  | 
2  | 
Author: Makarius  | 
|
3  | 
*)  | 
|
4  | 
||
| 58889 | 5  | 
section \<open>Standard ML within the Isabelle environment\<close>  | 
| 56276 | 6  | 
|
7  | 
theory Examples  | 
|
8  | 
imports Pure  | 
|
9  | 
begin  | 
|
10  | 
||
| 58616 | 11  | 
text \<open>  | 
| 61757 | 12  | 
Isabelle/ML is a somewhat augmented version of Standard ML, with various  | 
13  | 
add-ons that are indispensable for Isabelle development, but may cause  | 
|
14  | 
conflicts with independent projects in plain Standard ML.  | 
|
| 56276 | 15  | 
|
| 61757 | 16  | 
The Isabelle/Isar command \<^theory_text>\<open>SML_file\<close> supports official Standard ML within  | 
17  | 
the Isabelle environment, with full support in the Prover IDE  | 
|
| 56276 | 18  | 
(Isabelle/jEdit).  | 
19  | 
||
20  | 
Here is a very basic example that defines the factorial function and  | 
|
21  | 
evaluates it for some arguments.  | 
|
| 58616 | 22  | 
\<close>  | 
| 56276 | 23  | 
|
| 69605 | 24  | 
SML_file \<open>factorial.sml\<close>  | 
| 56276 | 25  | 
|
| 58616 | 26  | 
text \<open>  | 
| 61757 | 27  | 
The subsequent example illustrates the use of multiple \<^theory_text>\<open>SML_file\<close> commands  | 
28  | 
to build larger Standard ML projects. The toplevel SML environment is  | 
|
29  | 
enriched cumulatively within the theory context of Isabelle ---  | 
|
30  | 
independently of the Isabelle/ML environment.  | 
|
| 58616 | 31  | 
\<close>  | 
| 56276 | 32  | 
|
| 69605 | 33  | 
SML_file \<open>Example.sig\<close>  | 
34  | 
SML_file \<open>Example.sml\<close>  | 
|
| 56276 | 35  | 
|
| 
56618
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56276 
diff
changeset
 | 
36  | 
|
| 58616 | 37  | 
text \<open>  | 
| 
56618
 
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  | 
| 61757 | 39  | 
environments are separate. It is possible to exchange toplevel bindings via  | 
40  | 
separate commands like this.  | 
|
| 58616 | 41  | 
\<close>  | 
| 
56618
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56276 
diff
changeset
 | 
42  | 
|
| 61757 | 43  | 
SML_export \<open>val f = factorial\<close> \<comment> \<open>re-use factorial from SML environment\<close>  | 
| 58616 | 44  | 
ML \<open>f 42\<close>  | 
| 
56618
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56276 
diff
changeset
 | 
45  | 
|
| 58616 | 46  | 
SML_import \<open>val println = Output.writeln\<close>  | 
| 61757 | 47  | 
\<comment> \<open>re-use Isabelle/ML channel for SML\<close>  | 
| 
56618
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56276 
diff
changeset
 | 
48  | 
|
| 58616 | 49  | 
SML_import \<open>val par_map = Par_List.map\<close>  | 
| 61757 | 50  | 
\<comment> \<open>re-use Isabelle/ML parallel list combinator for SML\<close>  | 
| 
56618
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56276 
diff
changeset
 | 
51  | 
|
| 56276 | 52  | 
end  |