| author | Christian Urban <urbanc@in.tum.de> | 
| Wed, 20 Jun 2012 09:42:18 +0100 | |
| changeset 531 | cf7ef23348ed | 
| parent 517 | d8c376662bb4 | 
| child 542 | 4b96e3c8b33e | 
| permissions | -rw-r--r-- | 
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1 | (*<*) | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 2 | theory CallML | 
| 459 | 3 | imports "~~/src/HOL/Number_Theory/Primes" | 
| 4 | "~~/src/HOL/Library/Efficient_Nat" | |
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 5 | begin | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 6 | (*>*) | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 7 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 8 | section {* Calling ML Functions from within HOL \label{rec:callml} *}
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 9 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 10 | text{* 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 11 |   {\bf Problem:} 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 12 | How to call ML functions from within HOL?\smallskip | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 13 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 14 |   {\bf Solution:} This can be achieved with \isacommand{code\_const}
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 15 |   and \isacommand{code\_reflect}.\smallskip
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 16 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 17 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 18 | To make it clear we mean here calling unverified ML functions from within | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 19 |   HOL! The motivation is the paradigm of \emph{result checking}: rather than
 | 
| 441 | 20 | verifying some complicated algorithm, have the algorithm produce an easily | 
| 21 | checkable certificate. For example, instead of verifying an algorithm for | |
| 22 | testing non-primality, have an algorithm that produces a factor as a witness to | |
| 23 | non-primality. | |
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 24 | |
| 441 | 25 | The algorithm is an ML function finding a factor of a number. We first | 
| 26 | declare its type: | |
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 27 | *} | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 28 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 29 | consts factor :: "nat \<Rightarrow> nat" | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 30 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 31 | text{* 
 | 
| 441 | 32 | Its definition will be given below in ML. But the whole point is that | 
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 33 |   we can prove non-primality via @{const factor}, no matter what its
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 34 | actual definition is: | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 35 | *} | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 36 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 37 | lemma factor_non_prime: | 
| 428 | 38 | "let k = factor n in k \<noteq> 1 \<and> k \<noteq> n \<and> k dvd n \<Longrightarrow> \<not> prime n" | 
| 39 | by (auto simp: prime_nat_def Let_def) | |
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 40 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 41 | text{* 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 42 | Note that the premise is executable once we have defined | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 43 |   @{const factor}. Here is a trivial definition in ML: 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 44 | *} | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 45 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
459diff
changeset | 46 | ML %grayML{*fun factor n = if n = 4 then 2 else 1*}
 | 
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 47 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 48 | text{* 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 49 |   Of course this trivial definition of @{const factor} could have been given
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 50 | directly in HOL rather than ML. But by going to the ML level, all of ML is | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 51 | at our disposal, including arrays and references, features that are less | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 52 | easily emulated in HOL. In fact, we could even call some external software | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 53 | from ML, e.g.\ a computer algebra system. | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 54 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 55 | It should be noted, however, that in this example you need to import the | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 56 |   theory @{theory Efficient_Nat} in order to force the HOL-type @{typ nat} to
 | 
| 531 | 57 |   be implemented by the ML-type @{text "int"}. 
 | 
| 58 | ||
| 59 |   \begin{graybox}\small
 | |
| 60 |   \isacommand{theory}~@{text CallML}\\
 | |
| 61 |   \isacommand{imports}~@{text Main} @{text [quotes] "~~/src/HOL/Library/Efficient_Nat"}\\
 | |
| 62 | ... | |
| 63 |   \end{graybox}
 | |
| 64 | ||
| 65 | ||
| 66 | Thus the ML implementation of | |
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 67 |   @{const factor} must be and is of type @{text "int -> int"}. Now it is time
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 68 | to connect the two levels: | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 69 | *} | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 70 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 71 | code_const factor (SML "factor") | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 72 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 73 | text{* 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 74 |   The result of this declaration is that the HOL-function @{const factor} 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 75 | is executable and command | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 76 | *} | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 77 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 78 | value "factor 4" | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 79 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 80 | text{* 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 81 |   yields the expected result @{text 2}. Similarly we can prove that
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 82 |   @{text 4} is not prime: 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 83 | *} | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 84 | |
| 448 
957f69b9b7df
added something about Goal.prove_multi
 Christian Urban <urbanc@in.tum.de> parents: 
441diff
changeset | 85 | lemma "\<not> prime (4::nat)" | 
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 86 | apply(rule factor_non_prime) | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 87 | apply eval | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 88 | done | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 89 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 90 | text{* 
 | 
| 428 | 91 |   Note, however, the command \isacommand{code\_const} cannot check that the ML function
 | 
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 92 | has the required type. Therefore in the worst case a type mismatch will be detected by | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 93 |   the ML-compiler when we try to evaluate an expression involving @{const
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 94 |   factor}. It could also happen that @{const factor} is (accidentally)
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 95 | redefined on the ML level later on. But remember that we do not assume | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 96 |   anything about @{const factor} on the HOL-level. Hence no definition of
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 97 |   @{const factor} can ever lead to an incorrect proof. Of course ``wrong''
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 98 | definitions can lead to compile time or run time exceptions, or to failed | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 99 | proofs. | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 100 | |
| 428 | 101 | The above example was easy because we forced Isabelle (via the inclusion of the | 
| 102 |   theory @{theory
 | |
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 103 |   Efficient_Nat}) to implement @{typ nat} by @{text int}, a predefined 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 104 | ML-type. By default, Isabelle implements, for example, the HOL-type | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 105 |   @{text list} by the corresponding ML-type. Thus the following variation 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 106 |   on @{const factor} also works: 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 107 | *} | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 108 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 109 | consts factor2 :: "nat \<Rightarrow> nat list" (*<*)(*>*) | 
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
459diff
changeset | 110 | ML %grayML{*fun factor2 n = if n = 4 then [2] else []*}(*<*)(*>*)
 | 
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 111 | code_const factor2 (SML "factor2") | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 112 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 113 | value "factor2 4" | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 114 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 115 | text{* 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 116 |   The first line declares the type of @{const factor2}; the second
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 117 | gives its implementation in ML; the third makes it executable | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 118 | in HOL, and the last is just a test. In this way, you can easily | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 119 | interface with ML-functions whose types involve | 
| 428 | 120 |   @{text bool}, @{text int}, @{text list}, @{text option} and pairs, 
 | 
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 121 | only. If you have arbitrary tuples, for example, then you have to code | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 122 | them as nested pairs. | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 123 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 124 | Let us now look at how to refer to user-defined HOL-datatypes from the | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 125 |   ML-level. We modify our @{const factor} example a little by introducing a new
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 126 | datatype for the result: | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 127 | *} | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 128 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 129 | datatype result = Factor nat | Prime | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 130 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 131 | consts factor' :: "nat \<Rightarrow> result" | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 132 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 133 | text{* 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 134 | In order to write ML-code that uses this datatype, we need to define | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 135 | this datatype at the ML-level first. The following command does just that. | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 136 | *} | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 137 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 138 | code_reflect Result | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 139 | datatypes result = Factor | Prime | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 140 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 141 | text{* 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 142 |   This creates an ML-structure called @{text Result} (the name can be
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 143 |   arbitrarily chosen) that contains the datatype @{typ result}. The list 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 144 | of constructors (but not their types) needs to be given. Now we can | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 145 | write ML-code that uses this datatype: | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 146 | *} | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 147 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
459diff
changeset | 148 | ML %grayML{*fun factor' n = if n = 4 then Result.Factor 2 else Result.Prime*}
 | 
| 427 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 149 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 150 | text{* 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 151 |   Finally we can link the HOL and ML version of @{const factor'} as
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 152 | before: | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 153 | *} | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 154 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 155 | code_const factor' (SML "factor'") | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 156 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 157 | text{* 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 158 |   Now any evaluation of the HOL function @{const factor'} will use the
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 159 |   corresponding ML-function, like in the examples for @{const factor} above.
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 160 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 161 |   In general, \isacommand{code\_reflect} can export multiple datatypes 
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 162 |   (separated by \isacommand{and}) and also HOL-functions: simply add a line
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 163 |   \isacommand{functions} $f_1$ $f_2$ and so on.
 | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 164 | *} | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 165 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 166 | (*<*) | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 167 | value "factor' 4" | 
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 168 | |
| 
94538ddcac9b
added example from Tobias and changed the title
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 169 | end | 
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
459diff
changeset | 170 | (*>*) |