# HG changeset patch # User wenzelm # Date 1448889891 -3600 # Node ID 0d399131008f51dfcc49f1d698c3142e5c0d5987 # Parent 21c2b1f691d1756ec8cac37c0ba3e6380ae90bd8 tuned; diff -r 21c2b1f691d1 -r 0d399131008f src/HOL/ex/ML.thy --- a/src/HOL/ex/ML.thy Mon Nov 30 13:16:12 2015 +0100 +++ b/src/HOL/ex/ML.thy Mon Nov 30 14:24:51 2015 +0100 @@ -11,8 +11,8 @@ section \ML expressions\ text \ - The Isabelle command 'ML' allows to embed Isabelle/ML source into the formal - text. It is type-checked, compiled, and run within that environment. + The Isabelle command \<^theory_text>\ML\ allows to embed Isabelle/ML source into the + formal text. It is type-checked, compiled, and run within that environment. Note that side-effects should be avoided, unless the intention is to change global parameters of the run-time environment (rare). @@ -29,8 +29,10 @@ section \Antiquotations\ -text \There are some language extensions (via antiquotations), as explained in - the ``Isabelle/Isar implementation manual'', chapter 0.\ +text \ + There are some language extensions (via antiquotations), as explained in the + ``Isabelle/Isar implementation manual'', chapter 0. +\ ML \length []\ ML \@{assert} (length [] = 0)\ @@ -39,23 +41,33 @@ text \Formal entities from the surrounding context may be referenced as follows:\ -term "1 + 1" -- \term within theory source\ +term "1 + 1" \ \term within theory source\ ML \@{term "1 + 1"} (* term as symbolic ML datatype value *)\ ML \@{term "1 + (1::int)"}\ +ML \ + (* formal source with position information *) + val s = \1 + 1\; + + (* read term via old-style string interface *) + val t = Syntax.read_term @{context} (Syntax.implode_input s); +\ + section \IDE support\ text \ ML embedded into the Isabelle environment is connected to the Prover IDE. Poly/ML provides: - \begin{itemize} - \item precise positions for warnings / errors - \item markup for defining positions of identifiers - \item markup for inferred types of sub-expressions - \end{itemize} + + \<^item> precise positions for warnings / errors + \<^item> markup for defining positions of identifiers + \<^item> markup for inferred types of sub-expressions + \<^item> pretty-printing of ML values with markup + \<^item> completion of ML names + \<^item> source-level debugger \ ML \fn i => fn list => length list + i\ @@ -87,9 +99,10 @@ text \ Future.fork/join/cancel manage parallel evaluation. - Note that within Isabelle theory documents, the top-level command boundary may - not be transgressed without special precautions. This is normally managed by - the system when performing parallel proof checking.\ + Note that within Isabelle theory documents, the top-level command boundary + may not be transgressed without special precautions. This is normally + managed by the system when performing parallel proof checking. +\ ML \ val x = Future.fork (fn () => ackermann 3 10); @@ -97,8 +110,10 @@ val z = Future.join x + Future.join y \ -text \The @{ML_structure Par_List} module provides high-level combinators - for parallel list operations.\ +text \ + The @{ML_structure Par_List} module provides high-level combinators for + parallel list operations. +\ ML \timeit (fn () => map (fn n => ackermann 3 n) (1 upto 10))\ ML \timeit (fn () => Par_List.map (fn n => ackermann 3 n) (1 upto 10))\ @@ -111,12 +126,12 @@ "factorial 0 = 1" | "factorial (Suc n) = Suc n * factorial n" -term "factorial 4" -- \symbolic term\ -value "factorial 4" -- \evaluation via ML code generation in the background\ +term "factorial 4" \ \symbolic term\ +value "factorial 4" \ \evaluation via ML code generation in the background\ declare [[ML_source_trace]] -ML \@{term "factorial 4"}\ -- \symbolic term in ML\ -ML \@{code "factorial"}\ -- \ML code from function specification\ +ML \@{term "factorial 4"}\ \ \symbolic term in ML\ +ML \@{code "factorial"}\ \ \ML code from function specification\ fun ackermann :: "nat \ nat \ nat" diff -r 21c2b1f691d1 -r 0d399131008f src/Tools/SML/Examples.thy --- a/src/Tools/SML/Examples.thy Mon Nov 30 13:16:12 2015 +0100 +++ b/src/Tools/SML/Examples.thy Mon Nov 30 14:24:51 2015 +0100 @@ -9,12 +9,12 @@ begin text \ - Isabelle/ML is a somewhat augmented version of Standard ML, with - various add-ons that are indispensable for Isabelle development, but - may cause conflicts with independent projects in plain Standard ML. + Isabelle/ML is a somewhat augmented version of Standard ML, with various + add-ons that are indispensable for Isabelle development, but may cause + conflicts with independent projects in plain Standard ML. - The Isabelle/Isar command 'SML_file' supports official Standard ML - within the Isabelle environment, with full support in the Prover IDE + The Isabelle/Isar command \<^theory_text>\SML_file\ supports official Standard ML within + the Isabelle environment, with full support in the Prover IDE (Isabelle/jEdit). Here is a very basic example that defines the factorial function and @@ -24,10 +24,10 @@ SML_file "factorial.sml" text \ - The subsequent example illustrates the use of multiple 'SML_file' - commands to build larger Standard ML projects. The toplevel SML - environment is enriched cumulatively within the theory context of - Isabelle --- independently of the Isabelle/ML environment. + The subsequent example illustrates the use of multiple \<^theory_text>\SML_file\ commands + to build larger Standard ML projects. The toplevel SML environment is + enriched cumulatively within the theory context of Isabelle --- + independently of the Isabelle/ML environment. \ SML_file "Example.sig" @@ -36,17 +36,17 @@ text \ Isabelle/ML and SML share a common run-time system, but the static - environments are separate. It is possible to exchange toplevel bindings - via separate commands like this. + environments are separate. It is possible to exchange toplevel bindings via + separate commands like this. \ -SML_export \val f = factorial\ -- \re-use factorial from SML environment\ +SML_export \val f = factorial\ \ \re-use factorial from SML environment\ ML \f 42\ SML_import \val println = Output.writeln\ - -- \re-use Isabelle/ML channel for SML\ + \ \re-use Isabelle/ML channel for SML\ SML_import \val par_map = Par_List.map\ - -- \re-use Isabelle/ML parallel list combinator for SML\ + \ \re-use Isabelle/ML parallel list combinator for SML\ end