src/Doc/Codegen/Foundations.thy
author wenzelm
Tue, 28 Aug 2012 18:57:32 +0200
changeset 48985 5386df44a037
parent 48951 doc-src/Codegen/Foundations.thy@b9238cbcdd41
child 51143 0a2371e7ced3
permissions -rw-r--r--
renamed doc-src to src/Doc; renamed TutorialI to Tutorial;
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
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
     5
section {* Code generation foundations \label{sec:foundations} *}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
     6
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
     7
subsection {* Code generator architecture \label{sec:architecture} *}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
     8
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
     9
text {*
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]
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    22
    \includegraphics{architecture}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    23
    \caption{Code generator architecture}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    24
    \label{fig:arch}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    25
  \end{figure}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    26
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    27
  \noindent Central to code generation is the notion of \emph{code
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    28
  equations}.  A code equation as a first approximation is a theorem
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    29
  of the form @{text "f t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n \<equiv> t"} (an equation headed by a
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    30
  constant @{text f} with arguments @{text "t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n"} and right
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    31
  hand side @{text t}).
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    32
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    33
  \begin{itemize}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    34
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    35
    \item Starting point of code generation is a collection of (raw)
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    36
      code equations in a theory. It is not relevant where they stem
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    37
      from, but typically they were either produced by specification
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    38
      tools or proved explicitly by the user.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    39
      
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    40
    \item These raw code equations can be subjected to theorem
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    41
      transformations.  This \qn{preprocessor} (see
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    42
      \secref{sec:preproc}) can apply the full expressiveness of
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    43
      ML-based theorem transformations to code generation.  The result
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    44
      of preprocessing is a structured collection of code equations.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    45
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    46
    \item These code equations are \qn{translated} to a program in an
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    47
      abstract intermediate language.  Think of it as a kind of
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    48
      \qt{Mini-Haskell} with four \qn{statements}: @{text data} (for
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    49
      datatypes), @{text fun} (stemming from code equations), also
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    50
      @{text class} and @{text inst} (for type classes).
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    51
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    52
    \item Finally, the abstract program is \qn{serialised} into
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    53
      concrete source code of a target language.  This step only
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    54
      produces concrete syntax but does not change the program in
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    55
      essence; all conceptual transformations occur in the translation
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    56
      step.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    57
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    58
  \end{itemize}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    59
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    60
  \noindent From these steps, only the last two are carried out
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    61
  outside the logic; by keeping this layer as thin as possible, the
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    62
  amount of code to trust is kept to a minimum.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    63
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    64
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    65
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    66
subsection {* The preprocessor \label{sec:preproc} *}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    67
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    68
text {*
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    69
  Before selected function theorems are turned into abstract code, a
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    70
  chain of definitional transformation steps is carried out:
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    71
  \emph{preprocessing}.  The preprocessor consists of two
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    72
  components: a \emph{simpset} and \emph{function transformers}.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    73
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    74
  The \emph{simpset} can apply the full generality of the Isabelle
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    75
  simplifier.  Due to the interpretation of theorems as code
32000
6f07563dc8a9 updated to changes in sources; tuned
haftmann
parents: 31254
diff changeset
    76
  equations, rewrites are applied to the right hand side and the
6f07563dc8a9 updated to changes in sources; tuned
haftmann
parents: 31254
diff changeset
    77
  arguments of the left hand side of an equation, but never to the
6f07563dc8a9 updated to changes in sources; tuned
haftmann
parents: 31254
diff changeset
    78
  constant heading the left hand side.  An important special case are
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    79
  \emph{unfold theorems}, which may be declared and removed using the
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    80
  @{attribute code_unfold} or \emph{@{attribute code_unfold} del}
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
    81
  attribute, respectively.
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
    82
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    83
  Some common applications:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    84
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    85
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    86
text_raw {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    87
  \begin{itemize}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    88
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    89
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    90
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    91
     \item replacing non-executable constructs by executable ones:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    92
*}     
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    93
37211
32a6f471f090 clarified
haftmann
parents: 36259
diff changeset
    94
lemma %quote [code_unfold]:
37612
48fed6598be9 adapted to reorganization of auxiliary list operations; split off predicate compiler into separate theory
haftmann
parents: 37427
diff changeset
    95
  "x \<in> set xs \<longleftrightarrow> List.member xs x" by (fact in_set_member)
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    96
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    97
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    98
     \item replacing executable but inconvenient constructs:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    99
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   100
37211
32a6f471f090 clarified
haftmann
parents: 36259
diff changeset
   101
lemma %quote [code_unfold]:
37612
48fed6598be9 adapted to reorganization of auxiliary list operations; split off predicate compiler into separate theory
haftmann
parents: 37427
diff changeset
   102
  "xs = [] \<longleftrightarrow> List.null xs" by (fact eq_Nil_null)
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   103
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   104
text {*
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   105
     \item eliminating disturbing expressions:
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   106
*}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   107
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   108
lemma %quote [code_unfold]:
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   109
  "1 = Suc 0" by (fact One_nat_def)
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   110
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   111
text_raw {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   112
  \end{itemize}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   113
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   114
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   115
text {*
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   116
  \noindent \emph{Function transformers} provide a very general
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   117
  interface, transforming a list of function theorems to another list
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   118
  of function theorems, provided that neither the heading constant nor
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   119
  its type change.  The @{term "0\<Colon>nat"} / @{const Suc} pattern
46521
addcdf0dd283 clarified
haftmann
parents: 45211
diff changeset
   120
  elimination implemented in theory @{text Efficient_Nat} (see
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   121
  \secref{eff_nat}) uses this interface.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   122
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   123
  \noindent The current setup of the preprocessor may be inspected
38505
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   124
  using the @{command_def print_codeproc} command.  @{command_def
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   125
  code_thms} (see \secref{sec:equations}) provides a convenient
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   126
  mechanism to inspect the impact of a preprocessor setup on code
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   127
  equations.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   128
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   129
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   130
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   131
subsection {* Understanding code equations \label{sec:equations} *}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   132
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   133
text {*
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   134
  As told in \secref{sec:principle}, the notion of code equations is
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   135
  vital to code generation.  Indeed most problems which occur in
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   136
  practice can be resolved by an inspection of the underlying code
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   137
  equations.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   138
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   139
  It is possible to exchange the default code equations for constants
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   140
  by explicitly proving alternative ones:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   141
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   142
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   143
lemma %quote [code]:
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   144
  "dequeue (AQueue xs []) =
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   145
     (if xs = [] then (None, AQueue [] [])
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   146
       else dequeue (AQueue [] (rev xs)))"
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   147
  "dequeue (AQueue xs (y # ys)) =
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   148
     (Some y, AQueue xs ys)"
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   149
  by (cases xs, simp_all) (cases "rev xs", simp_all)
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   150
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   151
text {*
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   152
  \noindent The annotation @{text "[code]"} is an @{text attribute}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   153
  which states that the given theorems should be considered as code
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   154
  equations for a @{text fun} statement -- the corresponding constant
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   155
  is determined syntactically.  The resulting code:
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   156
*}
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   157
39745
3aa2bc9c5478 combine quote and typewriter/tt tag
haftmann
parents: 39683
diff changeset
   158
text %quotetypewriter {*
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39677
diff changeset
   159
  @{code_stmts dequeue (consts) dequeue (Haskell)}
39664
0afaf89ab591 more canonical type setting of type writer code examples
haftmann
parents: 38857
diff changeset
   160
*}
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   161
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   162
text {*
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   163
  \noindent You may note that the equality test @{term "xs = []"} has
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   164
  been replaced by the predicate @{term "List.null xs"}.  This is due
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   165
  to the default setup of the \qn{preprocessor}.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   166
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   167
  This possibility to select arbitrary code equations is the key
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   168
  technique for program and datatype refinement (see
39677
e9f89d86c963 corrected omission
haftmann
parents: 39664
diff changeset
   169
  \secref{sec:refinement}).
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   170
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   171
  Due to the preprocessor, there is the distinction of raw code
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   172
  equations (before preprocessing) and code equations (after
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   173
  preprocessing).
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   174
38505
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   175
  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
   176
  print_codesetup} command.
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   177
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   178
  The code equations after preprocessing are already are blueprint of
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   179
  the generated program and can be inspected using the @{command
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   180
  code_thms} command:
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   181
*}
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   182
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   183
code_thms %quote dequeue
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   184
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   185
text {*
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   186
  \noindent This prints a table with the code equations for @{const
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   187
  dequeue}, including \emph{all} code equations those equations depend
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   188
  on recursively.  These dependencies themselves can be visualized using
38505
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   189
  the @{command_def code_deps} command.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   190
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   191
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   192
30938
c6c9359e474c wellsortedness is no issue for a user manual any more
haftmann
parents: 30227
diff changeset
   193
subsection {* Equality *}
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   194
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   195
text {*
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   196
  Implementation of equality deserves some attention.  Here an example
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   197
  function involving polymorphic equality:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   198
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   199
28564
haftmann
parents: 28562
diff changeset
   200
primrec %quote collect_duplicates :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
28447
haftmann
parents: 28428
diff changeset
   201
  "collect_duplicates xs ys [] = xs"
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   202
| "collect_duplicates xs ys (z#zs) = (if z \<in> set xs
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   203
    then if z \<in> set ys
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   204
      then collect_duplicates xs ys zs
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   205
      else collect_duplicates xs (z#ys) zs
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   206
    else collect_duplicates (z#xs) (z#ys) zs)"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   207
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   208
text {*
37612
48fed6598be9 adapted to reorganization of auxiliary list operations; split off predicate compiler into separate theory
haftmann
parents: 37427
diff changeset
   209
  \noindent During preprocessing, the membership test is rewritten,
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   210
  resulting in @{const List.member}, which itself performs an explicit
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   211
  equality check, as can be seen in the corresponding @{text SML} code:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   212
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   213
39745
3aa2bc9c5478 combine quote and typewriter/tt tag
haftmann
parents: 39683
diff changeset
   214
text %quotetypewriter {*
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39677
diff changeset
   215
  @{code_stmts collect_duplicates (SML)}
39664
0afaf89ab591 more canonical type setting of type writer code examples
haftmann
parents: 38857
diff changeset
   216
*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   217
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   218
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   219
  \noindent Obviously, polymorphic equality is implemented the Haskell
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   220
  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
   221
  explicit class @{class equal} with a corresponding operation @{const
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38505
diff changeset
   222
  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
   223
  framework does the rest by propagating the @{class equal} constraints
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   224
  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
   225
  @{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
   226
  you may instantiate @{text equal} manually like any other type class.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   227
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   228
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   229
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   230
subsection {* Explicit partiality \label{sec:partiality} *}
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   231
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   232
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   233
  Partiality usually enters the game by partial patterns, as
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   234
  in the following example, again for amortised queues:
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   235
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   236
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   237
definition %quote strict_dequeue :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   238
  "strict_dequeue q = (case dequeue q
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   239
    of (Some x, q') \<Rightarrow> (x, q'))"
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   240
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   241
lemma %quote strict_dequeue_AQueue [code]:
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   242
  "strict_dequeue (AQueue xs (y # ys)) = (y, AQueue xs ys)"
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   243
  "strict_dequeue (AQueue xs []) =
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   244
    (case rev xs of y # ys \<Rightarrow> (y, AQueue [] ys))"
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   245
  by (simp_all add: strict_dequeue_def) (cases xs, simp_all split: list.split)
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   246
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   247
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   248
  \noindent In the corresponding code, there is no equation
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   249
  for the pattern @{term "AQueue [] []"}:
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   250
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   251
39745
3aa2bc9c5478 combine quote and typewriter/tt tag
haftmann
parents: 39683
diff changeset
   252
text %quotetypewriter {*
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39677
diff changeset
   253
  @{code_stmts strict_dequeue (consts) strict_dequeue (Haskell)}
39664
0afaf89ab591 more canonical type setting of type writer code examples
haftmann
parents: 38857
diff changeset
   254
*}
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   255
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   256
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   257
  \noindent In some cases it is desirable to have this
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   258
  pseudo-\qt{partiality} more explicitly, e.g.~as follows:
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   259
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   260
28564
haftmann
parents: 28562
diff changeset
   261
axiomatization %quote empty_queue :: 'a
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   262
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   263
definition %quote strict_dequeue' :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   264
  "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
   265
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   266
lemma %quote strict_dequeue'_AQueue [code]:
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   267
  "strict_dequeue' (AQueue xs []) = (if xs = [] then empty_queue
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   268
     else strict_dequeue' (AQueue [] (rev xs)))"
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   269
  "strict_dequeue' (AQueue xs (y # ys)) =
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   270
     (y, AQueue xs ys)"
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   271
  by (simp_all add: strict_dequeue'_def split: list.splits)
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   272
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   273
text {*
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   274
  Observe that on the right hand side of the definition of @{const
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   275
  "strict_dequeue'"}, the unspecified constant @{const empty_queue} occurs.
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   276
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   277
  Normally, if constants without any code equations occur in a
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   278
  program, the code generator complains (since in most cases this is
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   279
  indeed an error).  But such constants can also be thought
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   280
  of as function definitions which always fail,
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   281
  since there is never a successful pattern match on the left hand
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   282
  side.  In order to categorise a constant into that category
38505
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38458
diff changeset
   283
  explicitly, use @{command_def "code_abort"}:
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   284
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   285
28564
haftmann
parents: 28562
diff changeset
   286
code_abort %quote empty_queue
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   287
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   288
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   289
  \noindent Then the code generator will just insert an error or
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   290
  exception at the appropriate position:
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   291
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   292
39745
3aa2bc9c5478 combine quote and typewriter/tt tag
haftmann
parents: 39683
diff changeset
   293
text %quotetypewriter {*
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39677
diff changeset
   294
  @{code_stmts strict_dequeue' (consts) empty_queue strict_dequeue' (Haskell)}
39664
0afaf89ab591 more canonical type setting of type writer code examples
haftmann
parents: 38857
diff changeset
   295
*}
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   296
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   297
text {*
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   298
  \noindent This feature however is rarely needed in practice.  Note
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   299
  also that the HOL default setup already declares @{const undefined}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   300
  as @{command "code_abort"}, which is most likely to be used in such
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   301
  situations.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   302
*}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   303
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   304
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   305
subsection {* If something goes utterly wrong \label{sec:utterly_wrong} *}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   306
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   307
text {*
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   308
  Under certain circumstances, the code generator fails to produce
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   309
  code entirely.  To debug these, the following hints may prove
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   310
  helpful:
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   311
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   312
  \begin{description}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   313
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   314
    \ditem{\emph{Check with a different target language}.}  Sometimes
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   315
      the situation gets more clear if you switch to another target
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   316
      language; the code generated there might give some hints what
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   317
      prevents the code generator to produce code for the desired
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   318
      language.
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   319
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   320
    \ditem{\emph{Inspect code equations}.}  Code equations are the central
43410
957617fe0765 tuned spelling
haftmann
parents: 43407
diff changeset
   321
      carrier of code generation.  Most problems occurring while generating
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   322
      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
   323
      the error message.  A closer inspection of those may offer the key
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   324
      for solving issues (cf.~\secref{sec:equations}).
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   325
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   326
    \ditem{\emph{Inspect preprocessor setup}.}  The preprocessor might
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   327
      transform code equations unexpectedly; to understand an
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   328
      inspection of its setup is necessary (cf.~\secref{sec:preproc}).
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   329
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   330
    \ditem{\emph{Generate exceptions}.}  If the code generator
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   331
      complains about missing code equations, in can be helpful to
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   332
      implement the offending constants as exceptions
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   333
      (cf.~\secref{sec:partiality}); this allows at least for a formal
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   334
      generation of code, whose inspection may then give clues what is
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   335
      wrong.
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   336
38440
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   337
    \ditem{\emph{Remove offending code equations}.}  If code
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   338
      generation is prevented by just a single equation, this can be
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   339
      removed (cf.~\secref{sec:equations}) to allow formal code
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   340
      generation, whose result in turn can be used to trace the
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   341
      problem.  The most prominent case here are mismatches in type
6c0d02f416ba section "if something goes utterly wrong"
haftmann
parents: 38437
diff changeset
   342
      class signatures (\qt{wellsortedness error}).
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   343
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   344
  \end{description}
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   345
*}
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   346
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   347
end