src/Doc/Codegen/Evaluation.thy
author haftmann
Wed, 22 Feb 2017 20:24:50 +0100
changeset 65041 2525e680f94f
parent 63670 8e0148e1f5f4
child 66453 cc19f7ca2ed6
permissions -rw-r--r--
basic documentation for computations
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
     1
theory Evaluation
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
     2
imports Setup
59378
065f349852e6 separate image for prerequisites of codegen tutorial
haftmann
parents: 59377
diff changeset
     3
begin (*<*)
065f349852e6 separate image for prerequisites of codegen tutorial
haftmann
parents: 59377
diff changeset
     4
065f349852e6 separate image for prerequisites of codegen tutorial
haftmann
parents: 59377
diff changeset
     5
ML \<open>
065f349852e6 separate image for prerequisites of codegen tutorial
haftmann
parents: 59377
diff changeset
     6
  Isabelle_System.mkdirs (File.tmp_path (Path.basic "examples"))
065f349852e6 separate image for prerequisites of codegen tutorial
haftmann
parents: 59377
diff changeset
     7
\<close> (*>*)
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
     8
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
     9
section \<open>Evaluation \label{sec:evaluation}\<close>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    10
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    11
text \<open>
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    12
  Recalling \secref{sec:principle}, code generation turns a system of
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    13
  equations into a program with the \emph{same} equational semantics.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    14
  As a consequence, this program can be used as a \emph{rewrite
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    15
  engine} for terms: rewriting a term @{term "t"} using a program to a
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    16
  term @{term "t'"} yields the theorems @{prop "t \<equiv> t'"}.  This
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    17
  application of code generation in the following is referred to as
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    18
  \emph{evaluation}.
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    19
\<close>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    20
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    21
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    22
subsection \<open>Evaluation techniques\<close>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    23
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    24
text \<open>
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    25
  There is a rich palette of evaluation
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    26
  techniques, each comprising different aspects:
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    27
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    28
  \begin{description}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    29
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    30
    \item[Expressiveness.]  Depending on the extent to which symbolic
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    31
      computation is possible, the class of terms which can be evaluated
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    32
      can be bigger or smaller.
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    33
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    34
    \item[Efficiency.]  The more machine-near the technique, the
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    35
      faster it is.
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    36
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    37
    \item[Trustability.]  Techniques which a huge (and also probably
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    38
      more configurable infrastructure) are more fragile and less
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    39
      trustable.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    40
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    41
  \end{description}
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    42
\<close>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    43
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    44
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    45
subsubsection \<open>The simplifier (@{text simp})\<close>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    46
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    47
text \<open>
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    48
  The simplest way for evaluation is just using the simplifier with
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    49
  the original code equations of the underlying program.  This gives
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    50
  fully symbolic evaluation and highest trustablity, with the usual
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    51
  performance of the simplifier.  Note that for operations on abstract
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    52
  datatypes (cf.~\secref{sec:invariant}), the original theorems as
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    53
  given by the users are used, not the modified ones.
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    54
\<close>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    55
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    56
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    57
subsubsection \<open>Normalization by evaluation (@{text nbe})\<close>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    58
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    59
text \<open>
58620
7435b6a3f72e more antiquotations;
wenzelm
parents: 58310
diff changeset
    60
  Normalization by evaluation @{cite "Aehlig-Haftmann-Nipkow:2008:nbe"}
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    61
  provides a comparably fast partially symbolic evaluation which
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    62
  permits also normalization of functions and uninterpreted symbols;
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    63
  the stack of code to be trusted is considerable.
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    64
\<close>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    65
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    66
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    67
subsubsection \<open>Evaluation in ML (@{text code})\<close>
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    68
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    69
text \<open>
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    70
  Considerable performance can be achieved using evaluation in ML, at the cost
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    71
  of being restricted to ground results and a layered stack of code to
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    72
  be trusted, including a user's specific code generator setup.
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    73
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    74
  Evaluation is carried out in a target language \emph{Eval} which
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    75
  inherits from \emph{SML} but for convenience uses parts of the
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    76
  Isabelle runtime environment.  Hence soundness depends crucially
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    77
  on the correctness of the code generator setup; this is one of the
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    78
  reasons why you should not use adaptation (see \secref{sec:adaptation})
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    79
  frivolously.
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    80
\<close>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    81
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    82
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    83
subsection \<open>Dynamic evaluation\<close>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    84
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    85
text \<open>
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    86
  Dynamic evaluation takes the code generator configuration \qt{as it
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    87
  is} at the point where evaluation is issued and computes
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    88
  a corresponding result.  Best example is the
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
    89
  @{command_def value} command for ad-hoc evaluation of
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    90
  terms:
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    91
\<close>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    92
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    93
value %quote "42 / (12 :: rat)"
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    94
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
    95
text \<open>
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55418
diff changeset
    96
  \noindent @{command value} tries first to evaluate using ML, falling
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55418
diff changeset
    97
  back to normalization by evaluation if this fails.
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    98
58100
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
    99
  A particular technique may be specified in square brackets, e.g.
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
   100
\<close>
58100
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
   101
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
   102
value %quote [nbe] "42 / (12 :: rat)"
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
   103
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
   104
text \<open>
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   105
  To employ dynamic evaluation in documents, there is also
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55418
diff changeset
   106
  a @{text value} antiquotation with the same evaluation techniques 
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55418
diff changeset
   107
  as @{command value}.
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   108
\<close>
43656
9ece73262746 adding documentation of the value antiquotation to the code generation manual
bulwahn
parents: 41184
diff changeset
   109
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   110
  
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   111
subsubsection \<open>Term reconstruction in ML\<close>
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   112
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   113
text \<open>
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   114
  Results from evaluation in ML must be
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   115
  turned into Isabelle's internal term representation again.  Since
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   116
  that setup is highly configurable, it is never assumed to be trustable. 
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   117
  Hence evaluation in ML provides no full term reconstruction
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   118
  but distinguishes the following kinds:
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   119
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   120
  \begin{description}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   121
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   122
    \item[Plain evaluation.]  A term is normalized using the vanilla
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   123
      term reconstruction from ML to Isabelle; this is a pragmatic
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   124
      approach for applications which do not need trustability.
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   125
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   126
    \item[Property conversion.]  Evaluates propositions; since these
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   127
      are monomorphic, the term reconstruction is fixed once and for all
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   128
      and therefore trustable -- in the sense that only the regular
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   129
      code generator setup has to be trusted, without relying
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   130
      on term reconstruction from ML to Isabelle.
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   131
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   132
  \end{description}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   133
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   134
  \noindent The different degree of trustability is also manifest in the
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   135
  types of the corresponding ML functions: plain evaluation
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   136
  operates on uncertified terms, whereas property conversion
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   137
  operates on certified terms.
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
   138
\<close>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   139
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   140
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   141
subsubsection \<open>The partiality principle \label{sec:partiality_principle}\<close>
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   142
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   143
text \<open>
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   144
  During evaluation exceptions indicating a pattern
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   145
  match failure or a non-implemented function are treated specially:
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   146
  as sketched in \secref{sec:partiality}, such
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   147
  exceptions can be interpreted as partiality.  For plain evaluation,
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   148
  the result hence is optional; property conversion falls back
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   149
  to reflexivity in such cases.
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   150
\<close>
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   151
  
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   152
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   153
subsubsection \<open>Schematic overview\<close>
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   154
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
   155
text \<open>
39693
2ef15ec8e7dc tuned schema table
haftmann
parents: 39643
diff changeset
   156
  \newcommand{\ttsize}{\fontsize{5.8pt}{8pt}\selectfont}
2ef15ec8e7dc tuned schema table
haftmann
parents: 39643
diff changeset
   157
  \fontsize{9pt}{12pt}\selectfont
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   158
  \begin{tabular}{l||c|c|c}
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   159
    & @{text simp} & @{text nbe} & @{text code} \tabularnewline \hline \hline
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   160
    interactive evaluation & @{command value} @{text "[simp]"} & @{command value} @{text "[nbe]"} & @{command value} @{text "[code]"} \tabularnewline
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   161
    plain evaluation & & & \ttsize@{ML "Code_Evaluation.dynamic_value"} \tabularnewline \hline
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   162
    evaluation method & @{method code_simp} & @{method normalization} & @{method eval} \tabularnewline
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   163
    property conversion & & & \ttsize@{ML "Code_Runtime.dynamic_holds_conv"} \tabularnewline \hline
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   164
    conversion & \ttsize@{ML "Code_Simp.dynamic_conv"} & \ttsize@{ML "Nbe.dynamic_conv"}
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   165
  \end{tabular}
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
   166
\<close>
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   167
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   168
  
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   169
subsection \<open>Static evaluation\<close>
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   170
  
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
   171
text \<open>
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   172
  When implementing proof procedures using evaluation,
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   173
  in most cases the code generator setup is appropriate
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   174
  \qt{as it was} when the proof procedure was written in ML,
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   175
  not an arbitrary later potentially modified setup.  This is
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   176
  called static evaluation.
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
   177
\<close>
52287
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   178
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   179
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   180
subsubsection \<open>Static evaluation using @{text simp} and @{text nbe}\<close>
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   181
  
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
   182
text \<open>
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   183
  For @{text simp} and @{text nbe} static evaluation can be achieved using 
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   184
  @{ML Code_Simp.static_conv} and @{ML Nbe.static_conv}.
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   185
  Note that @{ML Nbe.static_conv} by its very nature
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   186
  requires an invocation of the ML compiler for every call,
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   187
  which can produce significant overhead.
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
   188
\<close>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   189
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   190
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   191
subsubsection \<open>Intimate connection between logic and system runtime\<close>
39609
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   192
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
   193
text \<open>
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   194
  Static evaluation for @{text eval} operates differently --
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   195
  the required generated code is inserted directly into an ML
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   196
  block using antiquotations.  The idea is that generated
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   197
  code performing static evaluation (called a \<^emph>\<open>computation\<close>)
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   198
  is compiled once and for all such that later calls do not
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   199
  require any invocation of the code generator or the ML
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   200
  compiler at all.  This topic deserved a dedicated chapter
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   201
  \secref{sec:computations}.
59377
056945909f60 modernized cartouches
haftmann
parents: 59335
diff changeset
   202
\<close>
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63670
diff changeset
   203
  
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   204
end