| author | wenzelm | 
| Sat, 10 Jan 2015 16:35:21 +0100 | |
| changeset 59342 | fd9102b419f5 | 
| parent 58889 | 5b7a9633cfa8 | 
| child 61757 | 0d399131008f | 
| 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> | 
| 56276 | 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. | |
| 58616 | 22 | \<close> | 
| 56276 | 23 | |
| 24 | SML_file "factorial.sml" | |
| 25 | ||
| 58616 | 26 | text \<open> | 
| 56276 | 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. | |
| 58616 | 31 | \<close> | 
| 56276 | 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: 
56276diff
changeset | 36 | |
| 58616 | 37 | text \<open> | 
| 56618 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 wenzelm parents: 
56276diff
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: 
56276diff
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: 
56276diff
changeset | 40 | via separate commands like this. | 
| 58616 | 41 | \<close> | 
| 56618 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 wenzelm parents: 
56276diff
changeset | 42 | |
| 58616 | 43 | SML_export \<open>val f = factorial\<close> -- \<open>re-use factorial from SML environment\<close> | 
| 44 | ML \<open>f 42\<close> | |
| 56618 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 wenzelm parents: 
56276diff
changeset | 45 | |
| 58616 | 46 | SML_import \<open>val println = Output.writeln\<close> | 
| 47 | -- \<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: 
56276diff
changeset | 48 | |
| 58616 | 49 | SML_import \<open>val par_map = Par_List.map\<close> | 
| 50 | -- \<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: 
56276diff
changeset | 51 | |
| 56276 | 52 | end |