src/Doc/Codegen/Evaluation.thy
author blanchet
Thu, 11 Sep 2014 18:54:36 +0200
changeset 58305 57752a91eec4
parent 58100 f54a8a4134d3
child 58310 91ea607a34d8
permissions -rw-r--r--
renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
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
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
     3
begin
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
     4
40751
6975c4d83ffd added label
haftmann
parents: 40350
diff changeset
     5
section {* Evaluation \label{sec:evaluation} *}
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
     6
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
     7
text {*
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
     8
  Recalling \secref{sec:principle}, code generation turns a system of
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
     9
  equations into a program with the \emph{same} equational semantics.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    10
  As a consequence, this program can be used as a \emph{rewrite
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    11
  engine} for terms: rewriting a term @{term "t"} using a program to a
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    12
  term @{term "t'"} yields the theorems @{prop "t \<equiv> t'"}.  This
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    13
  application of code generation in the following is referred to as
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    14
  \emph{evaluation}.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    15
*}
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    16
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    17
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    18
subsection {* Evaluation techniques *}
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    19
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    20
text {*
40350
1ef7ee8dd165 fixed typos
haftmann
parents: 39745
diff changeset
    21
  The existing infrastructure provides a rich palette of evaluation
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    22
  techniques, each comprising different aspects:
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    23
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    24
  \begin{description}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    25
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    26
    \item[Expressiveness.]  Depending on how good symbolic computation
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    27
      is supported, the class of terms which can be evaluated may be
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    28
      bigger or smaller.
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    29
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    30
    \item[Efficiency.]  The more machine-near the technique, the
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    31
      faster it is.
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    32
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    33
    \item[Trustability.]  Techniques which a huge (and also probably
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    34
      more configurable infrastructure) are more fragile and less
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    35
      trustable.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    36
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    37
  \end{description}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    38
*}
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    39
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    40
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    41
subsubsection {* The simplifier (@{text simp}) *}
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    42
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    43
text {*
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    44
  The simplest way for evaluation is just using the simplifier with
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    45
  the original code equations of the underlying program.  This gives
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    46
  fully symbolic evaluation and highest trustablity, with the usual
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    47
  performance of the simplifier.  Note that for operations on abstract
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    48
  datatypes (cf.~\secref{sec:invariant}), the original theorems as
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    49
  given by the users are used, not the modified ones.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    50
*}
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    51
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    52
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    53
subsubsection {* Normalization by evaluation (@{text nbe}) *}
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    54
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    55
text {*
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    56
  Normalization by evaluation \cite{Aehlig-Haftmann-Nipkow:2008:nbe}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    57
  provides a comparably fast partially symbolic evaluation which
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    58
  permits also normalization of functions and uninterpreted symbols;
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    59
  the stack of code to be trusted is considerable.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    60
*}
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    61
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    62
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    63
subsubsection {* Evaluation in ML (@{text code}) *}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    64
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    65
text {*
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    66
  Highest performance can be achieved by evaluation in ML, at the cost
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    67
  of being restricted to ground results and a layered stack of code to
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    68
  be trusted, including code generator configurations by the user.
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    69
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    70
  Evaluation is carried out in a target language \emph{Eval} which
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    71
  inherits from \emph{SML} but for convenience uses parts of the
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    72
  Isabelle runtime environment.  The soundness of computation carried
39609
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
    73
  out there depends crucially on the correctness of the code
39643
29cc021398fc corrections and tuning
haftmann
parents: 39609
diff changeset
    74
  generator setup; this is one of the reasons why you should not use
39609
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
    75
  adaptation (see \secref{sec:adaptation}) frivolously.
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    76
*}
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    77
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    78
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    79
subsection {* Aspects of evaluation *}
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    80
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    81
text {*
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    82
  Each of the techniques can be combined with different aspects.  The
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    83
  most important distinction is between dynamic and static evaluation.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    84
  Dynamic evaluation takes the code generator configuration \qt{as it
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    85
  is} at the point where evaluation is issued.  Best example is the
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    86
  @{command_def value} command which allows ad-hoc evaluation of
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
    87
  terms:
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    88
*}
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    89
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    90
value %quote "42 / (12 :: rat)"
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    91
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    92
text {*
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55418
diff changeset
    93
  \noindent @{command value} tries first to evaluate using ML, falling
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55418
diff changeset
    94
  back to normalization by evaluation if this fails.
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
    95
58100
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
    96
  A particular technique may be specified in square brackets, e.g.
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
    97
*}
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
    98
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
    99
value %quote [nbe] "42 / (12 :: rat)"
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
   100
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
   101
text {*
43656
9ece73262746 adding documentation of the value antiquotation to the code generation manual
bulwahn
parents: 41184
diff changeset
   102
  To employ dynamic evaluation in the document generation, there is also
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55418
diff changeset
   103
  a @{text value} antiquotation with the same evaluation techniques 
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55418
diff changeset
   104
  as @{command value}.
43656
9ece73262746 adding documentation of the value antiquotation to the code generation manual
bulwahn
parents: 41184
diff changeset
   105
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   106
  Static evaluation freezes the code generator configuration at a
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   107
  certain point and uses this context whenever evaluation is issued
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   108
  later on.  This is particularly appropriate for proof procedures
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   109
  which use evaluation, since then the behaviour of evaluation is not
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   110
  changed or even compromised later on by actions of the user.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   111
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   112
  As a technical complication, terms after evaluation in ML must be
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   113
  turned into Isabelle's internal term representation again.  Since
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   114
  this is also configurable, it is never fully trusted.  For this
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   115
  reason, evaluation in ML comes with further aspects:
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   116
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   117
  \begin{description}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   118
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   119
    \item[Plain evaluation.]  A term is normalized using the provided
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   120
      term reconstruction from ML to Isabelle; for applications which
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   121
      do not need to be fully trusted.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   122
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   123
    \item[Property conversion.]  Evaluates propositions; since these
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   124
      are monomorphic, the term reconstruction is fixed once and for all
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   125
      and therefore trustable.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   126
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   127
    \item[Conversion.]  Evaluates an arbitrary term @{term "t"} first
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   128
      by plain evaluation and certifies the result @{term "t'"} by
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   129
      checking the equation @{term "t \<equiv> t'"} using property
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   130
      conversion.
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
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   134
  \noindent The picture is further complicated by the roles of
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   135
  exceptions.  Here three cases have to be distinguished:
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   136
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   137
  \begin{itemize}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   138
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   139
    \item Evaluation of @{term t} terminates with a result @{term
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   140
      "t'"}.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   141
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   142
    \item Evaluation of @{term t} terminates which en exception
40350
1ef7ee8dd165 fixed typos
haftmann
parents: 39745
diff changeset
   143
      indicating a pattern match failure or a non-implemented
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   144
      function.  As sketched in \secref{sec:partiality}, this can be
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   145
      interpreted as partiality.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   146
     
39643
29cc021398fc corrections and tuning
haftmann
parents: 39609
diff changeset
   147
    \item Evaluation raises any other kind of exception.
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   148
     
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   149
  \end{itemize}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   150
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   151
  \noindent For conversions, the first case yields the equation @{term
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   152
  "t = t'"}, the second defaults to reflexivity @{term "t = t"}.
39643
29cc021398fc corrections and tuning
haftmann
parents: 39609
diff changeset
   153
  Exceptions of the third kind are propagated to the user.
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   154
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   155
  By default return values of plain evaluation are optional, yielding
40350
1ef7ee8dd165 fixed typos
haftmann
parents: 39745
diff changeset
   156
  @{text "SOME t'"} in the first case, @{text "NONE"} in the
1ef7ee8dd165 fixed typos
haftmann
parents: 39745
diff changeset
   157
  second, and propagating the exception in the third case.  A strict
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   158
  variant of plain evaluation either yields @{text "t'"} or propagates
51713
4fd969609b4d spelling
haftmann
parents: 48985
diff changeset
   159
  any exception, a liberal variant captures any exception in a result
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   160
  of type @{text "Exn.result"}.
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   161
  
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   162
  For property conversion (which coincides with conversion except for
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   163
  evaluation in ML), methods are provided which solve a given goal by
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   164
  evaluation.
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   165
*}
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   166
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   167
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   168
subsection {* Schematic overview *}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   169
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   170
text {*
39693
2ef15ec8e7dc tuned schema table
haftmann
parents: 39643
diff changeset
   171
  \newcommand{\ttsize}{\fontsize{5.8pt}{8pt}\selectfont}
2ef15ec8e7dc tuned schema table
haftmann
parents: 39643
diff changeset
   172
  \fontsize{9pt}{12pt}\selectfont
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   173
  \begin{tabular}{ll||c|c|c}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   174
    & & @{text simp} & @{text nbe} & @{text code} \tabularnewline \hline \hline
58100
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
   175
    \multirow{5}{1ex}{\rotatebox{90}{dynamic}}
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
   176
      & interactive evaluation 
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
   177
      & @{command value} @{text "[simp]"} & @{command value} @{text "[nbe]"} & @{command value} @{text "[code]"}
f54a8a4134d3 restored generic value slot, retaining default behaviour and separate approximate command
haftmann
parents: 56927
diff changeset
   178
      \tabularnewline
39693
2ef15ec8e7dc tuned schema table
haftmann
parents: 39643
diff changeset
   179
    & plain evaluation & & & \ttsize@{ML "Code_Evaluation.dynamic_value"} \tabularnewline \cline{2-5}
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   180
    & evaluation method & @{method code_simp} & @{method normalization} & @{method eval} \tabularnewline
39693
2ef15ec8e7dc tuned schema table
haftmann
parents: 39643
diff changeset
   181
    & property conversion & & & \ttsize@{ML "Code_Runtime.dynamic_holds_conv"} \tabularnewline \cline{2-5}
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40751
diff changeset
   182
    & conversion & \ttsize@{ML "Code_Simp.dynamic_conv"} & \ttsize@{ML "Nbe.dynamic_conv"}
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40751
diff changeset
   183
      & \ttsize@{ML "Code_Evaluation.dynamic_conv"} \tabularnewline \hline \hline
39693
2ef15ec8e7dc tuned schema table
haftmann
parents: 39643
diff changeset
   184
    \multirow{3}{1ex}{\rotatebox{90}{static}}
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40751
diff changeset
   185
    & plain evaluation & & & \ttsize@{ML "Code_Evaluation.static_value"} \tabularnewline \cline{2-5}
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   186
    & property conversion & &
39693
2ef15ec8e7dc tuned schema table
haftmann
parents: 39643
diff changeset
   187
      & \ttsize@{ML "Code_Runtime.static_holds_conv"} \tabularnewline \cline{2-5}
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40751
diff changeset
   188
    & conversion & \ttsize@{ML "Code_Simp.static_conv"}
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40751
diff changeset
   189
      & \ttsize@{ML "Nbe.static_conv"}
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40751
diff changeset
   190
      & \ttsize@{ML "Code_Evaluation.static_conv"}
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   191
  \end{tabular}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   192
*}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   193
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   194
52287
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   195
subsection {* Preprocessing HOL terms into evaluable shape *}
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   196
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   197
text {*
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   198
  When integration decision procedures developed inside HOL into HOL itself,
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   199
  it is necessary to somehow get from the Isabelle/ML representation to
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   200
  the representation used by the decision procedure itself (``reification'').
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   201
  One option is to hardcode it using code antiquotations (see \secref{sec:code_antiq}).
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   202
  Another option is to use pre-existing infrastructure in HOL:
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   203
  @{ML "Reification.conv"} and @{ML "Reification.tac"}
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   204
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   205
  An simplistic example:
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   206
*}
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   207
58305
57752a91eec4 renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents: 58100
diff changeset
   208
datatype_new %quote form_ord = T | F | Less nat nat
52287
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   209
  | And form_ord form_ord | Or form_ord form_ord | Neg form_ord
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   210
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   211
primrec %quote interp :: "form_ord \<Rightarrow> 'a::order list \<Rightarrow> bool"
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   212
where
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   213
  "interp T vs \<longleftrightarrow> True"
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   214
| "interp F vs \<longleftrightarrow> False"
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   215
| "interp (Less i j) vs \<longleftrightarrow> vs ! i < vs ! j"
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   216
| "interp (And f1 f2) vs \<longleftrightarrow> interp f1 vs \<and> interp f2 vs"
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   217
| "interp (Or f1 f2) vs \<longleftrightarrow> interp f1 vs \<or> interp f2 vs"
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   218
| "interp (Neg f) vs \<longleftrightarrow> \<not> interp f vs"
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   219
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   220
text {*
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   221
  The datatype @{type form_ord} represents formulae whose semantics is given by
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   222
  @{const interp}.  Note that values are represented by variable indices (@{typ nat})
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   223
  whose concrete values are given in list @{term vs}.
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   224
*}
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   225
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   226
ML (*<*) {* *}
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   227
schematic_lemma "thm": fixes x y z :: "'a::order" shows "x < y \<and> x < z \<equiv> ?P"
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   228
ML_prf 
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   229
(*>*) {* val thm =
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   230
  Reification.conv @{context} @{thms interp.simps} @{cterm "x < y \<and> x < z"} *} (*<*)
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   231
by (tactic {* ALLGOALS (rtac thm) *})
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   232
(*>*) 
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   233
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   234
text {*
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   235
  By virtue of @{fact interp.simps}, @{ML "Reification.conv"} provides a conversion
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   236
  which, for this concrete example, yields @{thm thm [no_vars]}.  Note that the argument
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   237
  to @{const interp} does not contain any free variables and can this be evaluated
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   238
  using evaluation.
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   239
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   240
  A less meager example can be found in the AFP, session @{text "Regular-Sets"},
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   241
  theory @{text Regexp_Method}.
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   242
*}
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   243
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   244
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   245
subsection {* Intimate connection between logic and system runtime *}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   246
39609
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   247
text {*
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   248
  The toolbox of static evaluation conversions forms a reasonable base
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   249
  to interweave generated code and system tools.  However in some
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   250
  situations more direct interaction is desirable.
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   251
*}
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   252
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   253
52287
7e54c4d964e7 some meagure hints concerning reification
haftmann
parents: 51713
diff changeset
   254
subsubsection {* Static embedding of generated code into system runtime -- the @{text code} antiquotation \label{sec:code_antiq} *}
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   255
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   256
text {*
39609
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   257
  The @{text code} antiquotation allows to include constants from
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   258
  generated code directly into ML system code, as in the following toy
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   259
  example:
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   260
*}
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   261
58305
57752a91eec4 renamed 'datatype' to 'old_datatype'; 'datatype' is now alias for 'datatype_new'
blanchet
parents: 58100
diff changeset
   262
datatype_new %quote form = T | F | And form form | Or form form (*<*)
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   263
39745
3aa2bc9c5478 combine quote and typewriter/tt tag
haftmann
parents: 39712
diff changeset
   264
(*>*) ML %quotett {*
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   265
  fun eval_form @{code T} = true
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   266
    | eval_form @{code F} = false
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   267
    | eval_form (@{code And} (p, q)) =
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   268
        eval_form p andalso eval_form q
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   269
    | eval_form (@{code Or} (p, q)) =
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   270
        eval_form p orelse eval_form q;
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   271
*}
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   272
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   273
text {*
39609
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   274
  \noindent @{text code} takes as argument the name of a constant;
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   275
  after the whole ML is read, the necessary code is generated
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   276
  transparently and the corresponding constant names are inserted.
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   277
  This technique also allows to use pattern matching on constructors
39643
29cc021398fc corrections and tuning
haftmann
parents: 39609
diff changeset
   278
  stemming from compiled datatypes.  Note that the @{text code}
29cc021398fc corrections and tuning
haftmann
parents: 39609
diff changeset
   279
  antiquotation may not refer to constants which carry adaptations;
29cc021398fc corrections and tuning
haftmann
parents: 39609
diff changeset
   280
  here you have to refer to the corresponding adapted code directly.
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   281
39609
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   282
  For a less simplistic example, theory @{text Approximation} in
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   283
  the @{text Decision_Procs} session is a good reference.
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   284
*}
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   285
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   286
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   287
subsubsection {* Static embedding of generated code into system runtime -- @{text code_reflect} *}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   288
39609
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   289
text {*
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   290
  The @{text code} antiquoation is lightweight, but the generated code
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   291
  is only accessible while the ML section is processed.  Sometimes this
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   292
  is not appropriate, especially if the generated code contains datatype
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   293
  declarations which are shared with other parts of the system.  In these
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   294
  cases, @{command_def code_reflect} can be used:
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   295
*}
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   296
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   297
code_reflect %quote Sum_Type
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   298
  datatypes sum = Inl | Inr
55418
9f25e0cca254 adapted to renaming of 'Projl' and 'Projr'
blanchet
parents: 52287
diff changeset
   299
  functions "Sum_Type.sum.projl" "Sum_Type.sum.projr"
39609
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   300
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   301
text {*
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   302
  \noindent @{command_def code_reflect} takes a structure name and
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   303
  references to datatypes and functions; for these code is compiled
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   304
  into the named ML structure and the \emph{Eval} target is modified
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   305
  in a way that future code generation will reference these
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   306
  precompiled versions of the given datatypes and functions.  This
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   307
  also allows to refer to the referenced datatypes and functions from
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   308
  arbitrary ML code as well.
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   309
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   310
  A typical example for @{command code_reflect} can be found in the
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   311
  @{theory Predicate} theory.
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   312
*}
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   313
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   314
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   315
subsubsection {* Separate compilation -- @{text code_reflect} *}
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   316
39609
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   317
text {*
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   318
  For technical reasons it is sometimes necessary to separate
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   319
  generation and compilation of code which is supposed to be used in
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   320
  the system runtime.  For this @{command code_reflect} with an
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   321
  optional @{text "file"} argument can be used:
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   322
*}
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   323
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   324
code_reflect %quote Rat
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   325
  datatypes rat = Frct
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   326
  functions Fract
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   327
    "(plus :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(minus :: rat \<Rightarrow> rat \<Rightarrow> rat)"
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   328
    "(times :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(divide :: rat \<Rightarrow> rat \<Rightarrow> rat)"
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   329
  file "examples/rat.ML"
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   330
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   331
text {*
39643
29cc021398fc corrections and tuning
haftmann
parents: 39609
diff changeset
   332
  \noindent This merely generates the referenced code to the given
39609
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   333
  file which can be included into the system runtime later on.
0af12dea761f sections on @{code} and code_reflect
haftmann
parents: 39599
diff changeset
   334
*}
39599
d9c247f7afa3 continued section abut evaluation
haftmann
parents: 39067
diff changeset
   335
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents:
diff changeset
   336
end
46522
2b1e87b3967f tuned whitespace
haftmann
parents: 43656
diff changeset
   337