src/Doc/Codegen/Foundations.thy
author blanchet
Tue, 03 Mar 2015 16:37:45 +0100
changeset 59574 de392405a851
parent 59377 056945909f60
child 61076 bdc1e2f0a86a
permissions -rw-r--r--
import 'Main' to be on the safe side
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
38405
7935b334893e sketch of new outline
haftmann
parents: 37612
diff changeset
     1
theory Foundations
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
     2
imports Introduction
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
     3
begin
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
     4
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
     5
section \<open>Code generation foundations \label{sec:foundations}\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
     6
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
     7
subsection \<open>Code generator architecture \label{sec:architecture}\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
     8
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
     9
text \<open>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    10
  The code generator is actually a framework consisting of different
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    11
  components which can be customised individually.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    12
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    13
  Conceptually all components operate on Isabelle's logic framework
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    14
  @{theory Pure}.  Practically, the object logic @{theory HOL}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    15
  provides the necessary facilities to make use of the code generator,
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    16
  mainly since it is an extension of @{theory Pure}.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    17
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    18
  The constellation of the different components is visualized in the
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    19
  following picture.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    20
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    21
  \begin{figure}[h]
52742
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    22
    \def\sys#1{\emph{#1}}
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    23
    \begin{tikzpicture}[x = 4cm, y = 1cm]
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    24
      \tikzstyle positive=[color = black, fill = white];
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    25
      \tikzstyle negative=[color = white, fill = black];
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    26
      \tikzstyle entity=[rounded corners, draw, thick];
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    27
      \tikzstyle process=[ellipse, draw, thick];
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    28
      \tikzstyle arrow=[-stealth, semithick];
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    29
      \node (spec) at (0, 3) [entity, positive] {specification tools};
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    30
      \node (user) at (1, 3) [entity, positive] {user proofs};
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    31
      \node (spec_user_join) at (0.5, 3) [shape=coordinate] {};
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    32
      \node (raw) at (0.5, 4) [entity, positive] {raw code equations};
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    33
      \node (pre) at (1.5, 4) [process, positive] {preprocessing};
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    34
      \node (eqn) at (2.5, 4) [entity, positive] {code equations};
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    35
      \node (iml) at (0.5, 0) [entity, positive] {intermediate program};
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    36
      \node (seri) at (1.5, 0) [process, positive] {serialisation};
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    37
      \node (SML) at (2.5, 3) [entity, positive] {\sys{SML}};
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    38
      \node (OCaml) at (2.5, 2) [entity, positive] {\sys{OCaml}};
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    39
      \node (Haskell) at (2.5, 1) [entity, positive] {\sys{Haskell}};
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    40
      \node (Scala) at (2.5, 0) [entity, positive] {\sys{Scala}};
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    41
      \draw [semithick] (spec) -- (spec_user_join);
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    42
      \draw [semithick] (user) -- (spec_user_join);
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    43
      \draw [-diamond, semithick] (spec_user_join) -- (raw);
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    44
      \draw [arrow] (raw) -- (pre);
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    45
      \draw [arrow] (pre) -- (eqn);
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    46
      \draw [arrow] (eqn) -- node (transl) [process, positive] {translation} (iml);
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    47
      \draw [arrow] (iml) -- (seri);
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    48
      \draw [arrow] (seri) -- (SML);
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    49
      \draw [arrow] (seri) -- (OCaml);
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    50
      \draw [arrow] (seri) -- (Haskell);
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    51
      \draw [arrow] (seri) -- (Scala);
e7296939fec2 more direct inclusion of tikz pictures;
wenzelm
parents: 51171
diff changeset
    52
    \end{tikzpicture}
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    53
    \caption{Code generator architecture}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    54
    \label{fig:arch}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    55
  \end{figure}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    56
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    57
  \noindent Central to code generation is the notion of \emph{code
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    58
  equations}.  A code equation as a first approximation is a theorem
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52742
diff changeset
    59
  of the form @{text "f t\<^sub>1 t\<^sub>2 \<dots> t\<^sub>n \<equiv> t"} (an equation headed by a
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52742
diff changeset
    60
  constant @{text f} with arguments @{text "t\<^sub>1 t\<^sub>2 \<dots> t\<^sub>n"} and right
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    61
  hand side @{text t}).
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    62
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    63
  \begin{itemize}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    64
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    65
    \item Starting point of code generation is a collection of (raw)
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    66
      code equations in a theory. It is not relevant where they stem
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    67
      from, but typically they were either produced by specification
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    68
      tools or proved explicitly by the user.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    69
      
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    70
    \item These raw code equations can be subjected to theorem
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    71
      transformations.  This \qn{preprocessor} (see
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    72
      \secref{sec:preproc}) can apply the full expressiveness of
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    73
      ML-based theorem transformations to code generation.  The result
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    74
      of preprocessing is a structured collection of code equations.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    75
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    76
    \item These code equations are \qn{translated} to a program in an
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    77
      abstract intermediate language.  Think of it as a kind of
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    78
      \qt{Mini-Haskell} with four \qn{statements}: @{text data} (for
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    79
      datatypes), @{text fun} (stemming from code equations), also
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    80
      @{text class} and @{text inst} (for type classes).
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    81
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    82
    \item Finally, the abstract program is \qn{serialised} into
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    83
      concrete source code of a target language.  This step only
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    84
      produces concrete syntax but does not change the program in
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    85
      essence; all conceptual transformations occur in the translation
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    86
      step.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    87
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    88
  \end{itemize}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    89
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    90
  \noindent From these steps, only the last two are carried out
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    91
  outside the logic; by keeping this layer as thin as possible, the
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    92
  amount of code to trust is kept to a minimum.
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
    93
\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    94
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    95
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
    96
subsection \<open>The pre- and postprocessor \label{sec:preproc}\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    97
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
    98
text \<open>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    99
  Before selected function theorems are turned into abstract code, a
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   100
  chain of definitional transformation steps is carried out:
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   101
  \emph{preprocessing}.  The preprocessor consists of two
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   102
  components: a \emph{simpset} and \emph{function transformers}.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   103
54721
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   104
  The preprocessor simpset has a disparate brother, the
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   105
  \emph{postprocessor simpset}.
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   106
  In the theory-to-code scenario depicted in the picture 
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   107
  above, it plays no role.  But if generated code is used
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   108
  to evaluate expressions (cf.~\secref{sec:evaluation}), the
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   109
  postprocessor simpset is applied to the resulting expression before this
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   110
  is turned back.
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   111
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   112
  The pre- and postprocessor \emph{simpsets} can apply the
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   113
  full generality of the Isabelle
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   114
  simplifier.  Due to the interpretation of theorems as code
32000
6f07563dc8a9 updated to changes in sources; tuned
haftmann
parents: 31254
diff changeset
   115
  equations, rewrites are applied to the right hand side and the
6f07563dc8a9 updated to changes in sources; tuned
haftmann
parents: 31254
diff changeset
   116
  arguments of the left hand side of an equation, but never to the
54721
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   117
  constant heading the left hand side.
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   118
54721
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   119
  Pre- and postprocessor can be setup to broker between
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   120
  expressions suitable for logical reasoning and expressions 
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   121
  suitable for execution.  As example, take list membership; logically
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   122
  is is just expressed as @{term "x \<in> set xs"}.  But for execution
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   123
  the intermediate set is not desirable.  Hence the following
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   124
  specification:
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   125
\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   126
54721
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   127
definition %quote member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool"
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   128
where
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   129
  [code_abbrev]: "member xs x \<longleftrightarrow> x \<in> set xs"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   130
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   131
text \<open>
54721
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   132
  \noindent The \emph{@{attribute code_abbrev} attribute} declares
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   133
  its theorem a rewrite rule for the postprocessor and the symmetric
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   134
  of its theorem as rewrite rule for the preprocessor.  Together,
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   135
  this has the effect that expressions @{thm (rhs) member_def [no_vars]}
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   136
  are replaced by @{thm (lhs) member_def [no_vars]} in generated code, but
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   137
  are turned back into @{thm (rhs) member_def [no_vars]} if generated
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   138
  code is used for evaluation.
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   139
54721
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   140
  Rewrite rules for pre- or postprocessor may be
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   141
  declared independently using \emph{@{attribute code_unfold}}
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   142
  or \emph{@{attribute code_post}} respectively.
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   143
54721
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   144
  \emph{Function transformers} provide a very general
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   145
  interface, transforming a list of function theorems to another list
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   146
  of function theorems, provided that neither the heading constant nor
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   147
  its type change.  The @{term "0\<Colon>nat"} / @{const Suc} pattern
51171
e8b2d90da499 corrected and clarified Code_Binary_Nat vs. Code_Target_Nat
haftmann
parents: 51143
diff changeset
   148
  used in theory @{text Code_Abstract_Nat} (see \secref{abstract_nat})
e8b2d90da499 corrected and clarified Code_Binary_Nat vs. Code_Target_Nat
haftmann
parents: 51143
diff changeset
   149
  uses this interface.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   150
54721
22b888402278 be more explicit about pre- and postprocessor, particularly code_abbrev
haftmann
parents: 53015
diff changeset
   151
  \noindent The current setup of the pre- and postprocessor may be inspected
38505
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   152
  using the @{command_def print_codeproc} command.  @{command_def
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   153
  code_thms} (see \secref{sec:equations}) provides a convenient
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   154
  mechanism to inspect the impact of a preprocessor setup on code
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   155
  equations.
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   156
\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   157
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   158
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   159
subsection \<open>Understanding code equations \label{sec:equations}\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   160
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   161
text \<open>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   162
  As told in \secref{sec:principle}, the notion of code equations is
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   163
  vital to code generation.  Indeed most problems which occur in
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   164
  practice can be resolved by an inspection of the underlying code
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   165
  equations.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   166
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   167
  It is possible to exchange the default code equations for constants
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   168
  by explicitly proving alternative ones:
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   169
\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   170
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   171
lemma %quote [code]:
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   172
  "dequeue (AQueue xs []) =
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   173
     (if xs = [] then (None, AQueue [] [])
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   174
       else dequeue (AQueue [] (rev xs)))"
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   175
  "dequeue (AQueue xs (y # ys)) =
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   176
     (Some y, AQueue xs ys)"
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   177
  by (cases xs, simp_all) (cases "rev xs", simp_all)
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   178
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   179
text \<open>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   180
  \noindent The annotation @{text "[code]"} is an @{text attribute}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   181
  which states that the given theorems should be considered as code
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   182
  equations for a @{text fun} statement -- the corresponding constant
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   183
  is determined syntactically.  The resulting code:
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   184
\<close>
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   185
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   186
text %quotetypewriter \<open>
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39677
diff changeset
   187
  @{code_stmts dequeue (consts) dequeue (Haskell)}
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   188
\<close>
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   189
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   190
text \<open>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   191
  \noindent You may note that the equality test @{term "xs = []"} has
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   192
  been replaced by the predicate @{term "List.null xs"}.  This is due
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   193
  to the default setup of the \qn{preprocessor}.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   194
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   195
  This possibility to select arbitrary code equations is the key
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   196
  technique for program and datatype refinement (see
39677
e9f89d86c963 corrected omission
haftmann
parents: 39664
diff changeset
   197
  \secref{sec:refinement}).
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   198
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   199
  Due to the preprocessor, there is the distinction of raw code
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   200
  equations (before preprocessing) and code equations (after
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   201
  preprocessing).
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   202
38505
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   203
  The first can be listed (among other data) using the @{command_def
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   204
  print_codesetup} command.
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   205
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   206
  The code equations after preprocessing are already are blueprint of
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   207
  the generated program and can be inspected using the @{command
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   208
  code_thms} command:
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   209
\<close>
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   210
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   211
code_thms %quote dequeue
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   212
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   213
text \<open>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   214
  \noindent This prints a table with the code equations for @{const
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   215
  dequeue}, including \emph{all} code equations those equations depend
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   216
  on recursively.  These dependencies themselves can be visualized using
38505
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   217
  the @{command_def code_deps} command.
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   218
\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   219
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   220
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   221
subsection \<open>Equality\<close>
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   222
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   223
text \<open>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   224
  Implementation of equality deserves some attention.  Here an example
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   225
  function involving polymorphic equality:
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   226
\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   227
28564
haftmann
parents: 28562
diff changeset
   228
primrec %quote collect_duplicates :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
28447
haftmann
parents: 28428
diff changeset
   229
  "collect_duplicates xs ys [] = xs"
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   230
| "collect_duplicates xs ys (z#zs) = (if z \<in> set xs
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   231
    then if z \<in> set ys
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   232
      then collect_duplicates xs ys zs
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   233
      else collect_duplicates xs (z#ys) zs
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   234
    else collect_duplicates (z#xs) (z#ys) zs)"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   235
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   236
text \<open>
37612
48fed6598be9 adapted to reorganization of auxiliary list operations; split off predicate compiler into separate theory
haftmann
parents: 37427
diff changeset
   237
  \noindent During preprocessing, the membership test is rewritten,
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   238
  resulting in @{const List.member}, which itself performs an explicit
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   239
  equality check, as can be seen in the corresponding @{text SML} code:
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   240
\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   241
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   242
text %quotetypewriter \<open>
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39677
diff changeset
   243
  @{code_stmts collect_duplicates (SML)}
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   244
\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   245
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   246
text \<open>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   247
  \noindent Obviously, polymorphic equality is implemented the Haskell
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   248
  way using a type class.  How is this achieved?  HOL introduces an
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38505
diff changeset
   249
  explicit class @{class equal} with a corresponding operation @{const
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38505
diff changeset
   250
  HOL.equal} such that @{thm equal [no_vars]}.  The preprocessing
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38505
diff changeset
   251
  framework does the rest by propagating the @{class equal} constraints
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   252
  through all dependent code equations.  For datatypes, instances of
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38505
diff changeset
   253
  @{class equal} are implicitly derived when possible.  For other types,
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38505
diff changeset
   254
  you may instantiate @{text equal} manually like any other type class.
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   255
\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   256
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   257
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   258
subsection \<open>Explicit partiality \label{sec:partiality}\<close>
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   259
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   260
text \<open>
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   261
  Partiality usually enters the game by partial patterns, as
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   262
  in the following example, again for amortised queues:
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   263
\<close>
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   264
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   265
definition %quote strict_dequeue :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   266
  "strict_dequeue q = (case dequeue q
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   267
    of (Some x, q') \<Rightarrow> (x, q'))"
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   268
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   269
lemma %quote strict_dequeue_AQueue [code]:
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   270
  "strict_dequeue (AQueue xs (y # ys)) = (y, AQueue xs ys)"
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   271
  "strict_dequeue (AQueue xs []) =
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   272
    (case rev xs of y # ys \<Rightarrow> (y, AQueue [] ys))"
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   273
  by (simp_all add: strict_dequeue_def) (cases xs, simp_all split: list.split)
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   274
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   275
text \<open>
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   276
  \noindent In the corresponding code, there is no equation
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   277
  for the pattern @{term "AQueue [] []"}:
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   278
\<close>
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   279
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   280
text %quotetypewriter \<open>
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39677
diff changeset
   281
  @{code_stmts strict_dequeue (consts) strict_dequeue (Haskell)}
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   282
\<close>
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   283
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   284
text \<open>
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   285
  \noindent In some cases it is desirable to have this
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   286
  pseudo-\qt{partiality} more explicitly, e.g.~as follows:
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   287
\<close>
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   288
28564
haftmann
parents: 28562
diff changeset
   289
axiomatization %quote empty_queue :: 'a
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   290
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   291
definition %quote strict_dequeue' :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   292
  "strict_dequeue' q = (case dequeue q of (Some x, q') \<Rightarrow> (x, q') | _ \<Rightarrow> empty_queue)"
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   293
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   294
lemma %quote strict_dequeue'_AQueue [code]:
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   295
  "strict_dequeue' (AQueue xs []) = (if xs = [] then empty_queue
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   296
     else strict_dequeue' (AQueue [] (rev xs)))"
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   297
  "strict_dequeue' (AQueue xs (y # ys)) =
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   298
     (y, AQueue xs ys)"
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   299
  by (simp_all add: strict_dequeue'_def split: list.splits)
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   300
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   301
text \<open>
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   302
  Observe that on the right hand side of the definition of @{const
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   303
  "strict_dequeue'"}, the unspecified constant @{const empty_queue} occurs.
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   304
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   305
  Normally, if constants without any code equations occur in a
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   306
  program, the code generator complains (since in most cases this is
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   307
  indeed an error).  But such constants can also be thought
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   308
  of as function definitions which always fail,
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   309
  since there is never a successful pattern match on the left hand
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   310
  side.  In order to categorise a constant into that category
54890
cb892d835803 fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents: 54721
diff changeset
   311
  explicitly, use the @{attribute code} attribute with
cb892d835803 fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents: 54721
diff changeset
   312
  @{text abort}:
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   313
\<close>
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   314
54890
cb892d835803 fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents: 54721
diff changeset
   315
declare %quote [[code abort: empty_queue]]
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   316
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   317
text \<open>
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   318
  \noindent Then the code generator will just insert an error or
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   319
  exception at the appropriate position:
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   320
\<close>
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   321
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   322
text %quotetypewriter \<open>
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39677
diff changeset
   323
  @{code_stmts strict_dequeue' (consts) empty_queue strict_dequeue' (Haskell)}
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   324
\<close>
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   325
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   326
text \<open>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   327
  \noindent This feature however is rarely needed in practice.  Note
54890
cb892d835803 fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents: 54721
diff changeset
   328
  also that the HOL default setup already declares @{const undefined},
cb892d835803 fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents: 54721
diff changeset
   329
  which is most likely to be used in such situations, as
cb892d835803 fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents: 54721
diff changeset
   330
  @{text "code abort"}.
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   331
\<close>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   332
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   333
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   334
subsection \<open>If something goes utterly wrong \label{sec:utterly_wrong}\<close>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   335
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   336
text \<open>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   337
  Under certain circumstances, the code generator fails to produce
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   338
  code entirely.  To debug these, the following hints may prove
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   339
  helpful:
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   340
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   341
  \begin{description}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   342
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   343
    \ditem{\emph{Check with a different target language}.}  Sometimes
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   344
      the situation gets more clear if you switch to another target
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   345
      language; the code generated there might give some hints what
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   346
      prevents the code generator to produce code for the desired
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   347
      language.
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   348
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   349
    \ditem{\emph{Inspect code equations}.}  Code equations are the central
43410
957617fe0765 tuned spelling
haftmann
parents: 43407
diff changeset
   350
      carrier of code generation.  Most problems occurring while generating
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   351
      code can be traced to single equations which are printed as part of
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   352
      the error message.  A closer inspection of those may offer the key
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   353
      for solving issues (cf.~\secref{sec:equations}).
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   354
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   355
    \ditem{\emph{Inspect preprocessor setup}.}  The preprocessor might
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   356
      transform code equations unexpectedly; to understand an
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   357
      inspection of its setup is necessary (cf.~\secref{sec:preproc}).
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   358
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   359
    \ditem{\emph{Generate exceptions}.}  If the code generator
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   360
      complains about missing code equations, in can be helpful to
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   361
      implement the offending constants as exceptions
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   362
      (cf.~\secref{sec:partiality}); this allows at least for a formal
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   363
      generation of code, whose inspection may then give clues what is
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   364
      wrong.
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   365
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   366
    \ditem{\emph{Remove offending code equations}.}  If code
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   367
      generation is prevented by just a single equation, this can be
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   368
      removed (cf.~\secref{sec:equations}) to allow formal code
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   369
      generation, whose result in turn can be used to trace the
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   370
      problem.  The most prominent case here are mismatches in type
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   371
      class signatures (\qt{wellsortedness error}).
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   372
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   373
  \end{description}
59377
056945909f60 modernized cartouches
haftmann
parents: 54890
diff changeset
   374
\<close>
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   375
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   376
end