| 38510 |      1 | theory Evaluation
 | 
|  |      2 | imports Setup
 | 
|  |      3 | begin
 | 
|  |      4 | 
 | 
|  |      5 | section {* Evaluation *}
 | 
|  |      6 | 
 | 
| 39599 |      7 | text {*
 | 
|  |      8 |   Recalling \secref{sec:principle}, code generation turns a system of
 | 
|  |      9 |   equations into a program with the \emph{same} equational semantics.
 | 
|  |     10 |   As a consequence, this program can be used as a \emph{rewrite
 | 
|  |     11 |   engine} for terms: rewriting a term @{term "t"} using a program to a
 | 
|  |     12 |   term @{term "t'"} yields the theorems @{prop "t \<equiv> t'"}.  This
 | 
|  |     13 |   application of code generation in the following is referred to as
 | 
|  |     14 |   \emph{evaluation}.
 | 
|  |     15 | *}
 | 
| 38510 |     16 | 
 | 
|  |     17 | 
 | 
|  |     18 | subsection {* Evaluation techniques *}
 | 
|  |     19 | 
 | 
| 39599 |     20 | text {*
 | 
|  |     21 |   The existing infrastructure provides a rich palett of evaluation
 | 
|  |     22 |   techniques, each comprising different aspects:
 | 
|  |     23 | 
 | 
|  |     24 |   \begin{description}
 | 
|  |     25 | 
 | 
|  |     26 |     \item[Expressiveness.]  Depending on how good symbolic computation
 | 
|  |     27 |       is supported, the class of terms which can be evaluated may be
 | 
|  |     28 |       bigger or smaller.
 | 
| 38510 |     29 | 
 | 
| 39599 |     30 |     \item[Efficiency.]  The more machine-near the technique, the
 | 
|  |     31 |       faster it is.
 | 
| 38510 |     32 | 
 | 
| 39599 |     33 |     \item[Trustability.]  Techniques which a huge (and also probably
 | 
|  |     34 |       more configurable infrastructure) are more fragile and less
 | 
|  |     35 |       trustable.
 | 
|  |     36 | 
 | 
|  |     37 |   \end{description}
 | 
|  |     38 | *}
 | 
| 38510 |     39 | 
 | 
|  |     40 | 
 | 
| 39599 |     41 | subsubsection {* The simplifier (@{text simp}) *}
 | 
| 38510 |     42 | 
 | 
| 39599 |     43 | text {*
 | 
|  |     44 |   The simplest way for evaluation is just using the simplifier with
 | 
|  |     45 |   the original code equations of the underlying program.  This gives
 | 
|  |     46 |   fully symbolic evaluation and highest trustablity, with the usual
 | 
|  |     47 |   performance of the simplifier.  Note that for operations on abstract
 | 
|  |     48 |   datatypes (cf.~\secref{sec:invariant}), the original theorems as
 | 
|  |     49 |   given by the users are used, not the modified ones.
 | 
|  |     50 | *}
 | 
| 38510 |     51 | 
 | 
|  |     52 | 
 | 
| 39599 |     53 | subsubsection {* Normalization by evaluation (@{text nbe}) *}
 | 
| 38510 |     54 | 
 | 
| 39599 |     55 | text {*
 | 
|  |     56 |   Normalization by evaluation \cite{Aehlig-Haftmann-Nipkow:2008:nbe}
 | 
|  |     57 |   provides a comparably fast partially symbolic evaluation which
 | 
|  |     58 |   permits also normalization of functions and uninterpreted symbols;
 | 
|  |     59 |   the stack of code to be trusted is considerable.
 | 
|  |     60 | *}
 | 
| 38510 |     61 | 
 | 
|  |     62 | 
 | 
| 39599 |     63 | subsubsection {* Evaluation in ML (@{text code}) *}
 | 
|  |     64 | 
 | 
|  |     65 | text {*
 | 
|  |     66 |   Highest performance can be achieved by evaluation in ML, at the cost
 | 
|  |     67 |   of being restricted to ground results and a layered stack of code to
 | 
|  |     68 |   be trusted, including code generator configurations by the user.
 | 
| 38510 |     69 | 
 | 
| 39599 |     70 |   Evaluation is carried out in a target language \emph{Eval} which
 | 
|  |     71 |   inherits from \emph{SML} but for convenience uses parts of the
 | 
|  |     72 |   Isabelle runtime environment.  The soundness of computation carried
 | 
| 39609 |     73 |   out there depends crucially on the correctness of the code
 | 
| 39643 |     74 |   generator setup; this is one of the reasons why you should not use
 | 
| 39609 |     75 |   adaptation (see \secref{sec:adaptation}) frivolously.
 | 
| 39599 |     76 | *}
 | 
| 38510 |     77 | 
 | 
|  |     78 | 
 | 
| 39599 |     79 | subsection {* Aspects of evaluation *}
 | 
| 38510 |     80 | 
 | 
|  |     81 | text {*
 | 
| 39599 |     82 |   Each of the techniques can be combined with different aspects.  The
 | 
|  |     83 |   most important distinction is between dynamic and static evaluation.
 | 
|  |     84 |   Dynamic evaluation takes the code generator configuration \qt{as it
 | 
|  |     85 |   is} at the point where evaluation is issued.  Best example is the
 | 
|  |     86 |   @{command_def value} command which allows ad-hoc evaluation of
 | 
|  |     87 |   terms:
 | 
| 38510 |     88 | *}
 | 
|  |     89 | 
 | 
|  |     90 | value %quote "42 / (12 :: rat)"
 | 
|  |     91 | 
 | 
|  |     92 | text {*
 | 
| 39599 |     93 |   \noindent By default @{command value} tries all available evaluation
 | 
|  |     94 |   techniques and prints the result of the first succeeding one.  A particular
 | 
|  |     95 |   technique may be specified in square brackets, e.g.
 | 
| 38510 |     96 | *}
 | 
|  |     97 | 
 | 
| 39599 |     98 | value %quote [nbe] "42 / (12 :: rat)"
 | 
| 38510 |     99 | 
 | 
|  |    100 | text {*
 | 
| 39599 |    101 |   Static evaluation freezes the code generator configuration at a
 | 
|  |    102 |   certain point and uses this context whenever evaluation is issued
 | 
|  |    103 |   later on.  This is particularly appropriate for proof procedures
 | 
|  |    104 |   which use evaluation, since then the behaviour of evaluation is not
 | 
|  |    105 |   changed or even compromised later on by actions of the user.
 | 
|  |    106 | 
 | 
|  |    107 |   As a technical complication, terms after evaluation in ML must be
 | 
|  |    108 |   turned into Isabelle's internal term representation again.  Since
 | 
|  |    109 |   this is also configurable, it is never fully trusted.  For this
 | 
|  |    110 |   reason, evaluation in ML comes with further aspects:
 | 
|  |    111 | 
 | 
|  |    112 |   \begin{description}
 | 
|  |    113 | 
 | 
|  |    114 |     \item[Plain evaluation.]  A term is normalized using the provided
 | 
|  |    115 |       term reconstruction from ML to Isabelle; for applications which
 | 
|  |    116 |       do not need to be fully trusted.
 | 
|  |    117 | 
 | 
|  |    118 |     \item[Property conversion.]  Evaluates propositions; since these
 | 
|  |    119 |       are monomorphic, the term reconstruction is fixed once and for all
 | 
|  |    120 |       and therefore trustable.
 | 
|  |    121 | 
 | 
|  |    122 |     \item[Conversion.]  Evaluates an arbitrary term @{term "t"} first
 | 
|  |    123 |       by plain evaluation and certifies the result @{term "t'"} by
 | 
|  |    124 |       checking the equation @{term "t \<equiv> t'"} using property
 | 
|  |    125 |       conversion.
 | 
|  |    126 | 
 | 
|  |    127 |   \end{description}
 | 
|  |    128 | 
 | 
|  |    129 |   \noindent The picture is further complicated by the roles of
 | 
|  |    130 |   exceptions.  Here three cases have to be distinguished:
 | 
|  |    131 | 
 | 
|  |    132 |   \begin{itemize}
 | 
|  |    133 | 
 | 
|  |    134 |     \item Evaluation of @{term t} terminates with a result @{term
 | 
|  |    135 |       "t'"}.
 | 
|  |    136 | 
 | 
|  |    137 |     \item Evaluation of @{term t} terminates which en exception
 | 
|  |    138 |       indicating a pattern match failure or a not-implemented
 | 
|  |    139 |       function.  As sketched in \secref{sec:partiality}, this can be
 | 
|  |    140 |       interpreted as partiality.
 | 
|  |    141 |      
 | 
| 39643 |    142 |     \item Evaluation raises any other kind of exception.
 | 
| 39599 |    143 |      
 | 
|  |    144 |   \end{itemize}
 | 
|  |    145 | 
 | 
|  |    146 |   \noindent For conversions, the first case yields the equation @{term
 | 
|  |    147 |   "t = t'"}, the second defaults to reflexivity @{term "t = t"}.
 | 
| 39643 |    148 |   Exceptions of the third kind are propagated to the user.
 | 
| 39599 |    149 | 
 | 
|  |    150 |   By default return values of plain evaluation are optional, yielding
 | 
|  |    151 |   @{text "SOME t'"} in the first case, @{text "NONE"} and in the
 | 
|  |    152 |   second propagating the exception in the third case.  A strict
 | 
|  |    153 |   variant of plain evaluation either yields @{text "t'"} or propagates
 | 
|  |    154 |   any exception, a liberal variant caputures any exception in a result
 | 
|  |    155 |   of type @{text "Exn.result"}.
 | 
|  |    156 |   
 | 
|  |    157 |   For property conversion (which coincides with conversion except for
 | 
|  |    158 |   evaluation in ML), methods are provided which solve a given goal by
 | 
|  |    159 |   evaluation.
 | 
| 38510 |    160 | *}
 | 
|  |    161 | 
 | 
|  |    162 | 
 | 
| 39599 |    163 | subsection {* Schematic overview *}
 | 
|  |    164 | 
 | 
|  |    165 | (*FIXME rotatebox?*)
 | 
| 38510 |    166 | 
 | 
|  |    167 | text {*
 | 
| 39599 |    168 |   \begin{tabular}{ll||c|c|c}
 | 
|  |    169 |     & & @{text simp} & @{text nbe} & @{text code} \tabularnewline \hline \hline
 | 
|  |    170 |     dynamic & interactive evaluation 
 | 
|  |    171 |       & @{command value} @{text "[simp]"} & @{command value} @{text "[nbe]"} & @{command value} @{text "[code]"}
 | 
|  |    172 |       \tabularnewline
 | 
|  |    173 |     & plain evaluation & & & @{ML "Code_Evaluation.dynamic_value"} \tabularnewline \hline
 | 
|  |    174 |     & evaluation method & @{method code_simp} & @{method normalization} & @{method eval} \tabularnewline
 | 
|  |    175 |     & property conversion & & & @{ML "Code_Runtime.dynamic_holds_conv"} \tabularnewline \hline
 | 
|  |    176 |     & conversion & @{ML "Code_Simp.dynamic_eval_conv"} & @{ML "Nbe.dynamic_eval_conv"}
 | 
|  |    177 |       & @{ML "Code_Evaluation.dynamic_eval_conv"} \tabularnewline \hline \hline
 | 
|  |    178 |     static & plain evaluation & & & @{ML "Code_Evaluation.static_value"} \tabularnewline \hline
 | 
|  |    179 |     & property conversion & &
 | 
|  |    180 |       & @{ML "Code_Runtime.static_holds_conv"} \tabularnewline \hline
 | 
|  |    181 |     & conversion & @{ML "Code_Simp.static_eval_conv"}
 | 
|  |    182 |       & @{ML "Nbe.static_eval_conv"}
 | 
|  |    183 |       & @{ML "Code_Evaluation.static_eval_conv"}
 | 
|  |    184 |   \end{tabular}
 | 
|  |    185 | *}
 | 
|  |    186 | 
 | 
|  |    187 | 
 | 
|  |    188 | subsection {* Intimate connection between logic and system runtime *}
 | 
|  |    189 | 
 | 
| 39609 |    190 | text {*
 | 
|  |    191 |   The toolbox of static evaluation conversions forms a reasonable base
 | 
|  |    192 |   to interweave generated code and system tools.  However in some
 | 
|  |    193 |   situations more direct interaction is desirable.
 | 
|  |    194 | *}
 | 
| 39599 |    195 | 
 | 
|  |    196 | 
 | 
| 39609 |    197 | subsubsection {* Static embedding of generated code into system runtime -- the @{text code} antiquotation *}
 | 
| 39599 |    198 | 
 | 
|  |    199 | text {*
 | 
| 39609 |    200 |   The @{text code} antiquotation allows to include constants from
 | 
|  |    201 |   generated code directly into ML system code, as in the following toy
 | 
|  |    202 |   example:
 | 
| 38510 |    203 | *}
 | 
|  |    204 | 
 | 
|  |    205 | datatype %quote form = T | F | And form form | Or form form (*<*)
 | 
|  |    206 | 
 | 
|  |    207 | (*>*) ML %quotett {*
 | 
|  |    208 |   fun eval_form @{code T} = true
 | 
|  |    209 |     | eval_form @{code F} = false
 | 
|  |    210 |     | eval_form (@{code And} (p, q)) =
 | 
|  |    211 |         eval_form p andalso eval_form q
 | 
|  |    212 |     | eval_form (@{code Or} (p, q)) =
 | 
|  |    213 |         eval_form p orelse eval_form q;
 | 
|  |    214 | *}
 | 
|  |    215 | 
 | 
|  |    216 | text {*
 | 
| 39609 |    217 |   \noindent @{text code} takes as argument the name of a constant;
 | 
|  |    218 |   after the whole ML is read, the necessary code is generated
 | 
|  |    219 |   transparently and the corresponding constant names are inserted.
 | 
|  |    220 |   This technique also allows to use pattern matching on constructors
 | 
| 39643 |    221 |   stemming from compiled datatypes.  Note that the @{text code}
 | 
|  |    222 |   antiquotation may not refer to constants which carry adaptations;
 | 
|  |    223 |   here you have to refer to the corresponding adapted code directly.
 | 
| 38510 |    224 | 
 | 
| 39609 |    225 |   For a less simplistic example, theory @{text Approximation} in
 | 
|  |    226 |   the @{text Decision_Procs} session is a good reference.
 | 
| 38510 |    227 | *}
 | 
|  |    228 | 
 | 
|  |    229 | 
 | 
| 39599 |    230 | subsubsection {* Static embedding of generated code into system runtime -- @{text code_reflect} *}
 | 
|  |    231 | 
 | 
| 39609 |    232 | text {*
 | 
|  |    233 |   The @{text code} antiquoation is lightweight, but the generated code
 | 
|  |    234 |   is only accessible while the ML section is processed.  Sometimes this
 | 
|  |    235 |   is not appropriate, especially if the generated code contains datatype
 | 
|  |    236 |   declarations which are shared with other parts of the system.  In these
 | 
|  |    237 |   cases, @{command_def code_reflect} can be used:
 | 
|  |    238 | *}
 | 
|  |    239 | 
 | 
|  |    240 | code_reflect %quote Sum_Type
 | 
|  |    241 |   datatypes sum = Inl | Inr
 | 
|  |    242 |   functions "Sum_Type.Projl" "Sum_Type.Projr"
 | 
|  |    243 | 
 | 
|  |    244 | text {*
 | 
|  |    245 |   \noindent @{command_def code_reflect} takes a structure name and
 | 
|  |    246 |   references to datatypes and functions; for these code is compiled
 | 
|  |    247 |   into the named ML structure and the \emph{Eval} target is modified
 | 
|  |    248 |   in a way that future code generation will reference these
 | 
|  |    249 |   precompiled versions of the given datatypes and functions.  This
 | 
|  |    250 |   also allows to refer to the referenced datatypes and functions from
 | 
|  |    251 |   arbitrary ML code as well.
 | 
|  |    252 | 
 | 
|  |    253 |   A typical example for @{command code_reflect} can be found in the
 | 
|  |    254 |   @{theory Predicate} theory.
 | 
|  |    255 | *}
 | 
|  |    256 | 
 | 
| 39599 |    257 | 
 | 
|  |    258 | subsubsection {* Separate compilation -- @{text code_reflect} *}
 | 
|  |    259 | 
 | 
| 39609 |    260 | text {*
 | 
|  |    261 |   For technical reasons it is sometimes necessary to separate
 | 
|  |    262 |   generation and compilation of code which is supposed to be used in
 | 
|  |    263 |   the system runtime.  For this @{command code_reflect} with an
 | 
|  |    264 |   optional @{text "file"} argument can be used:
 | 
|  |    265 | *}
 | 
|  |    266 | 
 | 
|  |    267 | code_reflect %quote Rat
 | 
|  |    268 |   datatypes rat = Frct
 | 
|  |    269 |   functions Fract
 | 
|  |    270 |     "(plus :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(minus :: rat \<Rightarrow> rat \<Rightarrow> rat)"
 | 
|  |    271 |     "(times :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(divide :: rat \<Rightarrow> rat \<Rightarrow> rat)"
 | 
|  |    272 |   file "examples/rat.ML"
 | 
|  |    273 | 
 | 
|  |    274 | text {*
 | 
| 39643 |    275 |   \noindent This merely generates the referenced code to the given
 | 
| 39609 |    276 |   file which can be included into the system runtime later on.
 | 
|  |    277 | *}
 | 
| 39599 |    278 | 
 | 
| 38510 |    279 | end
 |