ProgTutorial/Recipes/CallML.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 30 Apr 2012 14:43:52 +0100
changeset 517 d8c376662bb4
parent 459 4532577b61e0
child 531 cf7ef23348ed
permissions -rw-r--r--
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
4532577b61e0 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 448
diff changeset
     3
imports "~~/src/HOL/Number_Theory/Primes" 
4532577b61e0 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 448
diff changeset
     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
Christian Urban <urbanc@in.tum.de>
parents: 428
diff changeset
    20
  verifying some complicated algorithm, have the algorithm produce an easily
Christian Urban <urbanc@in.tum.de>
parents: 428
diff changeset
    21
  checkable certificate. For example, instead of verifying an algorithm for 
Christian Urban <urbanc@in.tum.de>
parents: 428
diff changeset
    22
  testing non-primality, have an algorithm that produces a factor as a witness to
Christian Urban <urbanc@in.tum.de>
parents: 428
diff changeset
    23
  non-primality. 
427
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
441
Christian Urban <urbanc@in.tum.de>
parents: 428
diff changeset
    25
  The algorithm is an ML function finding a factor of a number. We first 
Christian Urban <urbanc@in.tum.de>
parents: 428
diff changeset
    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
Christian Urban <urbanc@in.tum.de>
parents: 428
diff changeset
    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
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
    38
  "let k = factor n in k \<noteq> 1 \<and> k \<noteq> n \<and> k dvd n \<Longrightarrow> \<not> prime n"
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
    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: 459
diff 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
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
  be implemented by the ML-type @{text "int"}. Thus the ML implementation of
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
  @{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
    59
  to connect the two levels:
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
*}
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
code_const factor (SML "factor")
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  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
    66
  is executable and command 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
*}
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
value "factor 4"
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
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
  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
    73
  @{text 4} is not prime: 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
*}
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
448
957f69b9b7df added something about Goal.prove_multi
Christian Urban <urbanc@in.tum.de>
parents: 441
diff changeset
    76
lemma "\<not> prime (4::nat)"
427
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
  apply(rule factor_non_prime)
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  apply eval
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
  done
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
text{* 
428
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
    82
  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
    83
  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
    84
  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
    85
  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
    86
  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
    87
  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
    88
  @{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
    89
  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
    90
  proofs.
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
428
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
    92
  The above example was easy because we forced Isabelle (via the inclusion of the
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
    93
  theory @{theory
427
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
  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
    95
  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
    96
  @{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
    97
  on @{const factor} also works: 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
*}
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
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: 459
diff changeset
   101
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
   102
code_const factor2 (SML "factor2")
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
value "factor2 4"
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
  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
   108
  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
   109
  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
   110
  interface with ML-functions whose types involve 
428
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
   111
  @{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
   112
  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
   113
  them as nested pairs.
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
  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
   116
  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
   117
  datatype for the result: 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
*}
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
datatype result = Factor nat | Prime
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
consts factor' :: "nat \<Rightarrow> result"
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
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
  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
   126
  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
   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
code_reflect Result
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
  datatypes result = Factor | Prime
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  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
   134
  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
   135
  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
   136
  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
   137
*}
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 459
diff changeset
   139
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
   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
  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
   143
  before: 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
*}
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
code_const factor' (SML "factor'")
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
  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
   150
  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
   151
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
  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
   153
  (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
   154
  \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
   155
*}
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
(*<*)
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
value "factor' 4"
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
end
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 459
diff changeset
   161
(*>*)