ProgTutorial/Recipes/CallML.thy
author schropp <schropp@in.tum.de>
Mon, 31 May 2010 23:29:06 +0200
changeset 431 17f70e2834f5
parent 428 c2dde8c7174f
child 441 520127b708e6
permissions -rw-r--r--
added some further ho-pat-unif examples but commented out (missing response check)
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
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
imports "~~/src/HOL/Number_Theory/Primes" Efficient_Nat
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
begin
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
(*>*)
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
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
     8
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
  {\bf Problem:} 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
  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
    12
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
  {\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
    14
  and \isacommand{code\_reflect}.\smallskip
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
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
  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
    18
  HOL! The motivation is the paradigm of \emph{result checking}: rather than
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  verify some complicated algorithm, have the algorithm produce an easily
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  checkable certificate. For example, rather than verify a primality test,
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  have the test produce a factor of the input as a witness to
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  non-primality. This is only an example, and we ignore the converse
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  certificate, that for primality, although it, too, can be treated in the
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  same manner. 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
428
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
    26
  The example we are looking at here is an ML function for finding a factor 
427
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
  of a number. We first declare its type:
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
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
consts factor :: "nat \<Rightarrow> nat"
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
  Its definition will be given in ML below. But the whole point is that
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
  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
    35
  actual definition is: 
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
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
lemma factor_non_prime:
428
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
    39
  "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
    40
  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
    41
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
  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
    44
  @{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
    45
*}
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
ML{*fun factor n = if n = 4 then 2 else 1*}
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  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
    51
  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
    52
  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
    53
  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
    54
  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
    55
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  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
    57
  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
    58
  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
    59
  @{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
    60
  to connect the two levels:
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
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
code_const factor (SML "factor")
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
  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
    67
  is executable and command 
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
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
value "factor 4"
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
  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
    74
  @{text 4} is not prime: 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
*}
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
lemma "\<not> prime(4::nat)"
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  apply(rule factor_non_prime)
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
  apply eval
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
  done
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
text{* 
428
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
    83
  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
    84
  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
    85
  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
    86
  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
    87
  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
    88
  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
    89
  @{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
    90
  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
    91
  proofs.
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
428
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
    93
  The above example was easy because we forced Isabelle (via the inclusion of the
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
    94
  theory @{theory
427
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
  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
    96
  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
    97
  @{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
    98
  on @{const factor} also works: 
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
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
consts factor2 :: "nat \<Rightarrow> nat list" (*<*)(*>*)
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
ML{*fun factor2 n = if n = 4 then [2] else []*}(*<*)(*>*)
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
code_const factor2 (SML "factor2")
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
value "factor2 4"
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
  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
   109
  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
   110
  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
   111
  interface with ML-functions whose types involve 
428
Christian Urban <urbanc@in.tum.de>
parents: 427
diff changeset
   112
  @{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
   113
  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
   114
  them as nested pairs.
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
  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
   117
  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
   118
  datatype for the result: 
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
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
datatype result = Factor nat | Prime
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
consts factor' :: "nat \<Rightarrow> result"
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
  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
   127
  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
   128
*}
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
code_reflect Result
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  datatypes result = Factor | Prime
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
  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
   135
  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
   136
  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
   137
  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
   138
*}
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
ML{*fun factor' n = if n = 4 then Result.Factor 2 else Result.Prime*}
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
  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
   144
  before: 
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
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
code_const factor' (SML "factor'")
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
text{* 
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
  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
   151
  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
   152
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
  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
   154
  (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
   155
  \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
   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
(*<*)
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
value "factor' 4"
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
end
94538ddcac9b added example from Tobias and changed the title
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
(*>*)