doc-src/IsarAdvanced/Codegen/Thy/Program.thy
author haftmann
Tue, 03 Feb 2009 19:37:16 +0100
changeset 29794 32d00a2a6f28
parent 29560 fa6c5d62adf5
child 29796 a342da8ddf39
permissions -rw-r--r--
added stub about datatype abstraction
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
     1
theory Program
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
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
     5
section {* Turning Theories into Programs \label{sec:program} *}
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
     6
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
     7
subsection {* The @{text "Isabelle/HOL"} default setup *}
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
     8
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
     9
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    10
  We have already seen how by default equations stemming from
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    11
  @{command definition}/@{command primrec}/@{command fun}
28447
haftmann
parents: 28428
diff changeset
    12
  statements are used for code generation.  This default behaviour
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
    13
  can be changed, e.g. by providing different code equations.
28593
f087237af65d continued codegen tutorial
haftmann
parents: 28564
diff changeset
    14
  All kinds of customisation shown in this section is \emph{safe}
28447
haftmann
parents: 28428
diff changeset
    15
  in the sense that the user does not have to worry about
haftmann
parents: 28428
diff changeset
    16
  correctness -- all programs generatable that way are partially
haftmann
parents: 28428
diff changeset
    17
  correct.
haftmann
parents: 28428
diff changeset
    18
*}
haftmann
parents: 28428
diff changeset
    19
haftmann
parents: 28428
diff changeset
    20
subsection {* Selecting code equations *}
haftmann
parents: 28428
diff changeset
    21
haftmann
parents: 28428
diff changeset
    22
text {*
haftmann
parents: 28428
diff changeset
    23
  Coming back to our introductory example, we
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
    24
  could provide an alternative code equations for @{const dequeue}
28447
haftmann
parents: 28428
diff changeset
    25
  explicitly:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    26
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    27
28564
haftmann
parents: 28562
diff changeset
    28
lemma %quote [code]:
28447
haftmann
parents: 28428
diff changeset
    29
  "dequeue (Queue xs []) =
28456
haftmann
parents: 28447
diff changeset
    30
     (if xs = [] then (None, Queue [] [])
haftmann
parents: 28447
diff changeset
    31
       else dequeue (Queue [] (rev xs)))"
28447
haftmann
parents: 28428
diff changeset
    32
  "dequeue (Queue xs (y # ys)) =
haftmann
parents: 28428
diff changeset
    33
     (Some y, Queue xs ys)"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    34
  by (cases xs, simp_all) (cases "rev xs", simp_all)
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    35
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    36
text {*
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28462
diff changeset
    37
  \noindent The annotation @{text "[code]"} is an @{text Isar}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    38
  @{text attribute} which states that the given theorems should be
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
    39
  considered as code equations for a @{text fun} statement --
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    40
  the corresponding constant is determined syntactically.  The resulting code:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    41
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    42
28564
haftmann
parents: 28562
diff changeset
    43
text %quote {*@{code_stmts dequeue (consts) dequeue (Haskell)}*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    44
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    45
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    46
  \noindent You may note that the equality test @{term "xs = []"} has been
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    47
  replaced by the predicate @{term "null xs"}.  This is due to the default
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    48
  setup in the \qn{preprocessor} to be discussed further below (\secref{sec:preproc}).
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    49
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    50
  Changing the default constructor set of datatypes is also
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
    51
  possible.  See \secref{sec:datatypes} for an example.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    52
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    53
  As told in \secref{sec:concept}, code generation is based
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    54
  on a structured collection of code theorems.
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    55
  For explorative purpose, this collection
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    56
  may be inspected using the @{command code_thms} command:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    57
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    58
28564
haftmann
parents: 28562
diff changeset
    59
code_thms %quote dequeue
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    60
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    61
text {*
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
    62
  \noindent prints a table with \emph{all} code equations
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    63
  for @{const dequeue}, including
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
    64
  \emph{all} code equations those equations depend
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    65
  on recursively.
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    66
  
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    67
  Similarly, the @{command code_deps} command shows a graph
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
    68
  visualising dependencies between code equations.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    69
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    70
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    71
subsection {* @{text class} and @{text instantiation} *}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    72
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    73
text {*
28447
haftmann
parents: 28428
diff changeset
    74
  Concerning type classes and code generation, let us examine an example
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    75
  from abstract algebra:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    76
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    77
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
    78
class %quote semigroup =
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    79
  fixes mult :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<otimes>" 70)
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    80
  assumes assoc: "(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    81
28564
haftmann
parents: 28562
diff changeset
    82
class %quote monoid = semigroup +
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    83
  fixes neutral :: 'a ("\<one>")
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    84
  assumes neutl: "\<one> \<otimes> x = x"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    85
    and neutr: "x \<otimes> \<one> = x"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    86
28564
haftmann
parents: 28562
diff changeset
    87
instantiation %quote nat :: monoid
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    88
begin
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    89
28564
haftmann
parents: 28562
diff changeset
    90
primrec %quote mult_nat where
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    91
    "0 \<otimes> n = (0\<Colon>nat)"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    92
  | "Suc m \<otimes> n = n + m \<otimes> n"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    93
28564
haftmann
parents: 28562
diff changeset
    94
definition %quote neutral_nat where
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    95
  "\<one> = Suc 0"
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
    96
28564
haftmann
parents: 28562
diff changeset
    97
lemma %quote add_mult_distrib:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    98
  fixes n m q :: nat
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    99
  shows "(n + m) \<otimes> q = n \<otimes> q + m \<otimes> q"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   100
  by (induct n) simp_all
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   101
28564
haftmann
parents: 28562
diff changeset
   102
instance %quote proof
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   103
  fix m n q :: nat
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   104
  show "m \<otimes> n \<otimes> q = m \<otimes> (n \<otimes> q)"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   105
    by (induct m) (simp_all add: add_mult_distrib)
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   106
  show "\<one> \<otimes> n = n"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   107
    by (simp add: neutral_nat_def)
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   108
  show "m \<otimes> \<one> = m"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   109
    by (induct m) (simp_all add: neutral_nat_def)
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   110
qed
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   111
28564
haftmann
parents: 28562
diff changeset
   112
end %quote
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   113
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   114
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   115
  \noindent We define the natural operation of the natural numbers
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   116
  on monoids:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   117
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   118
28564
haftmann
parents: 28562
diff changeset
   119
primrec %quote (in monoid) pow :: "nat \<Rightarrow> 'a \<Rightarrow> 'a" where
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   120
    "pow 0 a = \<one>"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   121
  | "pow (Suc n) a = a \<otimes> pow n a"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   122
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   123
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   124
  \noindent This we use to define the discrete exponentiation function:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   125
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   126
28564
haftmann
parents: 28562
diff changeset
   127
definition %quote bexp :: "nat \<Rightarrow> nat" where
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   128
  "bexp n = pow n (Suc (Suc 0))"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   129
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   130
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   131
  \noindent The corresponding code:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   132
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   133
28564
haftmann
parents: 28562
diff changeset
   134
text %quote {*@{code_stmts bexp (Haskell)}*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   135
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   136
text {*
28447
haftmann
parents: 28428
diff changeset
   137
  \noindent This is a convenient place to show how explicit dictionary construction
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   138
  manifests in generated code (here, the same example in @{text SML}):
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   139
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   140
28564
haftmann
parents: 28562
diff changeset
   141
text %quote {*@{code_stmts bexp (SML)}*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   142
28447
haftmann
parents: 28428
diff changeset
   143
text {*
haftmann
parents: 28428
diff changeset
   144
  \noindent Note the parameters with trailing underscore (@{verbatim "A_"})
haftmann
parents: 28428
diff changeset
   145
    which are the dictionary parameters.
haftmann
parents: 28428
diff changeset
   146
*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   147
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   148
subsection {* The preprocessor \label{sec:preproc} *}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   149
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   150
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   151
  Before selected function theorems are turned into abstract
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   152
  code, a chain of definitional transformation steps is carried
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   153
  out: \emph{preprocessing}.  In essence, the preprocessor
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   154
  consists of two components: a \emph{simpset} and \emph{function transformers}.
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   155
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   156
  The \emph{simpset} allows to employ the full generality of the Isabelle
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   157
  simplifier.  Due to the interpretation of theorems
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
   158
  as code equations, rewrites are applied to the right
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   159
  hand side and the arguments of the left hand side of an
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   160
  equation, but never to the constant heading the left hand side.
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   161
  An important special case are \emph{inline theorems} which may be
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   162
  declared and undeclared using the
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   163
  \emph{code inline} or \emph{code inline del} attribute respectively.
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   164
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   165
  Some common applications:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   166
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   167
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   168
text_raw {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   169
  \begin{itemize}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   170
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   171
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   172
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   173
     \item replacing non-executable constructs by executable ones:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   174
*}     
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   175
28564
haftmann
parents: 28562
diff changeset
   176
lemma %quote [code inline]:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   177
  "x \<in> set xs \<longleftrightarrow> x mem xs" by (induct xs) simp_all
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   178
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   179
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   180
     \item eliminating superfluous constants:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   181
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   182
28564
haftmann
parents: 28562
diff changeset
   183
lemma %quote [code inline]:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   184
  "1 = Suc 0" by simp
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   185
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   186
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   187
     \item replacing executable but inconvenient constructs:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   188
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   189
28564
haftmann
parents: 28562
diff changeset
   190
lemma %quote [code inline]:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   191
  "xs = [] \<longleftrightarrow> List.null xs" by (induct xs) simp_all
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   192
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   193
text_raw {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   194
  \end{itemize}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   195
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   196
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   197
text {*
28447
haftmann
parents: 28428
diff changeset
   198
  \noindent \emph{Function transformers} provide a very general interface,
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   199
  transforming a list of function theorems to another
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   200
  list of function theorems, provided that neither the heading
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   201
  constant nor its type change.  The @{term "0\<Colon>nat"} / @{const Suc}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   202
  pattern elimination implemented in
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   203
  theory @{text Efficient_Nat} (see \secref{eff_nat}) uses this
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   204
  interface.
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   205
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   206
  \noindent The current setup of the preprocessor may be inspected using
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   207
  the @{command print_codesetup} command.
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   208
  @{command code_thms} provides a convenient
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   209
  mechanism to inspect the impact of a preprocessor setup
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
   210
  on code equations.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   211
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   212
  \begin{warn}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   213
    The attribute \emph{code unfold}
28447
haftmann
parents: 28428
diff changeset
   214
    associated with the @{text "SML code generator"} also applies to
haftmann
parents: 28428
diff changeset
   215
    the @{text "generic code generator"}:
haftmann
parents: 28428
diff changeset
   216
    \emph{code unfold} implies \emph{code inline}.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   217
  \end{warn}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   218
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   219
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   220
subsection {* Datatypes \label{sec:datatypes} *}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   221
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   222
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   223
  Conceptually, any datatype is spanned by a set of
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   224
  \emph{constructors} of type @{text "\<tau> = \<dots> \<Rightarrow> \<kappa> \<alpha>\<^isub>1 \<dots> \<alpha>\<^isub>n"} where @{text
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   225
  "{\<alpha>\<^isub>1, \<dots>, \<alpha>\<^isub>n}"} is exactly the set of \emph{all} type variables in
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   226
  @{text "\<tau>"}.  The HOL datatype package by default registers any new
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   227
  datatype in the table of datatypes, which may be inspected using the
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   228
  @{command print_codesetup} command.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   229
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   230
  In some cases, it is appropriate to alter or extend this table.  As
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   231
  an example, we will develop an alternative representation of the
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   232
  queue example given in \secref{sec:intro}.  The amortised
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   233
  representation is convenient for generating code but exposes its
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   234
  \qt{implementation} details, which may be cumbersome when proving
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   235
  theorems about it.  Therefore, here a simple, straightforward
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   236
  representation of queues:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   237
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   238
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   239
datatype %quote 'a queue = Queue "'a list"
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   240
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   241
definition %quote empty :: "'a queue" where
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   242
  "empty = Queue []"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   243
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   244
primrec %quote enqueue :: "'a \<Rightarrow> 'a queue \<Rightarrow> 'a queue" where
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   245
  "enqueue x (Queue xs) = Queue (xs @ [x])"
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   246
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   247
fun %quote dequeue :: "'a queue \<Rightarrow> 'a option \<times> 'a queue" where
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   248
    "dequeue (Queue []) = (None, Queue [])"
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   249
  | "dequeue (Queue (x # xs)) = (Some x, Queue xs)"
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   250
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   251
text {*
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   252
  \noindent This we can use directly for proving;  for executing,
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   253
  we provide an alternative characterisation:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   254
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   255
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   256
definition %quote AQueue :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a queue" where
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   257
  "AQueue xs ys = Queue (ys @ rev xs)"
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   258
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   259
code_datatype %quote AQueue
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   260
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   261
text {* *}
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   262
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   263
lemma %quote empty_AQueue [code]:
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   264
  "empty = AQueue [] []"
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   265
  unfolding AQueue_def empty_def by simp
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   266
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   267
lemma %quote enqueue_AQueue [code]:
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   268
  "enqueue x (AQueue xs ys) = AQueue (x # xs) ys"
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   269
  unfolding AQueue_def by simp
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   270
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   271
lemma %quote dequeue_AQueue [code]:
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   272
  "dequeue (AQueue xs []) =
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   273
    (if xs = [] then (None, AQueue [] []) else dequeue (AQueue [] (rev xs)))"
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   274
  "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)"
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   275
  unfolding AQueue_def by simp_all
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   276
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   277
text {* *}
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   278
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   279
definition %quote aqueue_case [code inline]:
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   280
  "aqueue_case = queue_case"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   281
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   282
lemma %quote case_AQueue1 [code]:
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   283
  "queue_case f (AQueue xs ys) = f (ys @ rev xs)"
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   284
  unfolding AQueue_def by simp
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   285
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   286
lemma %quote case_AQueue2 [code]:
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   287
  "aqueue_case f (AQueue xs ys) = f (ys @ rev xs)"
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   288
  unfolding aqueue_case case_AQueue1 ..
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   289
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   290
text {* *}
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   291
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   292
text %quote {*@{code_stmts empty enqueue dequeue (SML)}*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   293
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   294
text {*
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   295
  \noindent From this example, it can be glimpsed that using own
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   296
  constructor sets is a little delicate since it changes the set of
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   297
  valid patterns for values of that type.  Without going into much
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   298
  detail, here some practical hints:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   299
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   300
  \begin{itemize}
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   301
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   302
    \item When changing the constructor set for datatypes, take care
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   303
      to provide an alternative for the @{text case} combinator
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   304
      (e.g.~by replacing it using the preprocessor).
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   305
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   306
    \item Values in the target language need not to be normalised --
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   307
      different values in the target language may represent the same
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   308
      value in the logic.
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   309
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   310
    \item Usually, a good methodology to deal with the subtleties of
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   311
      pattern matching is to see the type as an abstract type: provide
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   312
      a set of operations which operate on the concrete representation
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   313
      of the type, and derive further operations by combinations of
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   314
      these primitive ones, without relying on a particular
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   315
      representation.
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   316
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   317
  \end{itemize}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   318
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   319
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   320
code_datatype %invisible "0::nat" Suc
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28462
diff changeset
   321
declare %invisible plus_Dig [code del]
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   322
declare %invisible One_nat_def [code inline]
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28462
diff changeset
   323
declare %invisible add_Suc_shift [code] 
4e74209f113e `code func` now just `code`
haftmann
parents: 28462
diff changeset
   324
lemma %invisible [code]: "0 + n = (n \<Colon> nat)" by simp
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   325
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   326
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   327
subsection {* Equality and wellsortedness *}
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   328
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   329
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   330
  Surely you have already noticed how equality is treated
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   331
  by the code generator:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   332
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   333
28564
haftmann
parents: 28562
diff changeset
   334
primrec %quote collect_duplicates :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
28447
haftmann
parents: 28428
diff changeset
   335
  "collect_duplicates xs ys [] = xs"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   336
  | "collect_duplicates xs ys (z#zs) = (if z \<in> set xs
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   337
      then if z \<in> set ys
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   338
        then collect_duplicates xs ys zs
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   339
        else collect_duplicates xs (z#ys) zs
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   340
      else collect_duplicates (z#xs) (z#ys) zs)"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   341
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   342
text {*
28428
haftmann
parents: 28419
diff changeset
   343
  \noindent The membership test during preprocessing is rewritten,
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   344
  resulting in @{const List.member}, which itself
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   345
  performs an explicit equality check.
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   346
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   347
28564
haftmann
parents: 28562
diff changeset
   348
text %quote {*@{code_stmts collect_duplicates (SML)}*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   349
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   350
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   351
  \noindent Obviously, polymorphic equality is implemented the Haskell
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   352
  way using a type class.  How is this achieved?  HOL introduces
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   353
  an explicit class @{class eq} with a corresponding operation
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   354
  @{const eq_class.eq} such that @{thm eq [no_vars]}.
28447
haftmann
parents: 28428
diff changeset
   355
  The preprocessing framework does the rest by propagating the
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
   356
  @{class eq} constraints through all dependent code equations.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   357
  For datatypes, instances of @{class eq} are implicitly derived
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   358
  when possible.  For other types, you may instantiate @{text eq}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   359
  manually like any other type class.
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   360
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   361
  Though this @{text eq} class is designed to get rarely in
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   362
  the way, a subtlety
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   363
  enters the stage when definitions of overloaded constants
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   364
  are dependent on operational equality.  For example, let
28447
haftmann
parents: 28428
diff changeset
   365
  us define a lexicographic ordering on tuples
haftmann
parents: 28428
diff changeset
   366
  (also see theory @{theory Product_ord}):
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   367
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   368
28564
haftmann
parents: 28562
diff changeset
   369
instantiation %quote "*" :: (order, order) order
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   370
begin
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   371
28564
haftmann
parents: 28562
diff changeset
   372
definition %quote [code del]:
28447
haftmann
parents: 28428
diff changeset
   373
  "x \<le> y \<longleftrightarrow> fst x < fst y \<or> fst x = fst y \<and> snd x \<le> snd y"
28428
haftmann
parents: 28419
diff changeset
   374
28564
haftmann
parents: 28562
diff changeset
   375
definition %quote [code del]:
28447
haftmann
parents: 28428
diff changeset
   376
  "x < y \<longleftrightarrow> fst x < fst y \<or> fst x = fst y \<and> snd x < snd y"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   377
28564
haftmann
parents: 28562
diff changeset
   378
instance %quote proof
28447
haftmann
parents: 28428
diff changeset
   379
qed (auto simp: less_eq_prod_def less_prod_def intro: order_less_trans)
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   380
28564
haftmann
parents: 28562
diff changeset
   381
end %quote
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   382
28564
haftmann
parents: 28562
diff changeset
   383
lemma %quote order_prod [code]:
28447
haftmann
parents: 28428
diff changeset
   384
  "(x1 \<Colon> 'a\<Colon>order, y1 \<Colon> 'b\<Colon>order) < (x2, y2) \<longleftrightarrow>
haftmann
parents: 28428
diff changeset
   385
     x1 < x2 \<or> x1 = x2 \<and> y1 < y2"
haftmann
parents: 28428
diff changeset
   386
  "(x1 \<Colon> 'a\<Colon>order, y1 \<Colon> 'b\<Colon>order) \<le> (x2, y2) \<longleftrightarrow>
haftmann
parents: 28428
diff changeset
   387
     x1 < x2 \<or> x1 = x2 \<and> y1 \<le> y2"
haftmann
parents: 28428
diff changeset
   388
  by (simp_all add: less_prod_def less_eq_prod_def)
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   389
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   390
text {*
28428
haftmann
parents: 28419
diff changeset
   391
  \noindent Then code generation will fail.  Why?  The definition
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   392
  of @{term "op \<le>"} depends on equality on both arguments,
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   393
  which are polymorphic and impose an additional @{class eq}
28447
haftmann
parents: 28428
diff changeset
   394
  class constraint, which the preprocessor does not propagate
haftmann
parents: 28428
diff changeset
   395
  (for technical reasons).
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   396
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   397
  The solution is to add @{class eq} explicitly to the first sort arguments in the
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   398
  code theorems:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   399
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   400
28564
haftmann
parents: 28562
diff changeset
   401
lemma %quote order_prod_code [code]:
28447
haftmann
parents: 28428
diff changeset
   402
  "(x1 \<Colon> 'a\<Colon>{order, eq}, y1 \<Colon> 'b\<Colon>order) < (x2, y2) \<longleftrightarrow>
haftmann
parents: 28428
diff changeset
   403
     x1 < x2 \<or> x1 = x2 \<and> y1 < y2"
haftmann
parents: 28428
diff changeset
   404
  "(x1 \<Colon> 'a\<Colon>{order, eq}, y1 \<Colon> 'b\<Colon>order) \<le> (x2, y2) \<longleftrightarrow>
haftmann
parents: 28428
diff changeset
   405
     x1 < x2 \<or> x1 = x2 \<and> y1 \<le> y2"
haftmann
parents: 28428
diff changeset
   406
  by (simp_all add: less_prod_def less_eq_prod_def)
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   407
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   408
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   409
  \noindent Then code generation succeeds:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   410
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   411
28564
haftmann
parents: 28562
diff changeset
   412
text %quote {*@{code_stmts "op \<le> \<Colon> _ \<times> _ \<Rightarrow> _ \<times> _ \<Rightarrow> bool" (SML)}*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   413
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   414
text {*
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
   415
  In some cases, the automatically derived code equations
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   416
  for equality on a particular type may not be appropriate.
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   417
  As example, watch the following datatype representing
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   418
  monomorphic parametric types (where type constructors
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   419
  are referred to by natural numbers):
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   420
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   421
28564
haftmann
parents: 28562
diff changeset
   422
datatype %quote monotype = Mono nat "monotype list"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   423
(*<*)
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   424
lemma monotype_eq:
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   425
  "eq_class.eq (Mono tyco1 typargs1) (Mono tyco2 typargs2) \<equiv> 
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   426
     eq_class.eq tyco1 tyco2 \<and> eq_class.eq typargs1 typargs2" by (simp add: eq)
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   427
(*>*)
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   428
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   429
text {*
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   430
  \noindent Then code generation for SML would fail with a message
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   431
  that the generated code contains illegal mutual dependencies:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   432
  the theorem @{thm monotype_eq [no_vars]} already requires the
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   433
  instance @{text "monotype \<Colon> eq"}, which itself requires
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   434
  @{thm monotype_eq [no_vars]};  Haskell has no problem with mutually
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   435
  recursive @{text instance} and @{text function} definitions,
28593
f087237af65d continued codegen tutorial
haftmann
parents: 28564
diff changeset
   436
  but the SML serialiser does not support this.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   437
28447
haftmann
parents: 28428
diff changeset
   438
  In such cases, you have to provide your own equality equations
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   439
  involving auxiliary constants.  In our case,
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   440
  @{const [show_types] list_all2} can do the job:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   441
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   442
28564
haftmann
parents: 28562
diff changeset
   443
lemma %quote monotype_eq_list_all2 [code]:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   444
  "eq_class.eq (Mono tyco1 typargs1) (Mono tyco2 typargs2) \<longleftrightarrow>
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   445
     eq_class.eq tyco1 tyco2 \<and> list_all2 eq_class.eq typargs1 typargs2"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   446
  by (simp add: eq list_all2_eq [symmetric])
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   447
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   448
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   449
  \noindent does not depend on instance @{text "monotype \<Colon> eq"}:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   450
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   451
28564
haftmann
parents: 28562
diff changeset
   452
text %quote {*@{code_stmts "eq_class.eq :: monotype \<Rightarrow> monotype \<Rightarrow> bool" (SML)}*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   453
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   454
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   455
subsection {* Explicit partiality *}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   456
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   457
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   458
  Partiality usually enters the game by partial patterns, as
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   459
  in the following example, again for amortised queues:
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   460
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   461
28564
haftmann
parents: 28562
diff changeset
   462
fun %quote strict_dequeue :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   463
  "strict_dequeue (Queue xs (y # ys)) = (y, Queue xs ys)"
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   464
  | "strict_dequeue (Queue xs []) =
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   465
      (case rev xs of y # ys \<Rightarrow> (y, Queue [] ys))"
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   466
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   467
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   468
  \noindent In the corresponding code, there is no equation
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   469
  for the pattern @{term "Queue [] []"}:
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   470
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   471
28564
haftmann
parents: 28562
diff changeset
   472
text %quote {*@{code_stmts strict_dequeue (consts) strict_dequeue (Haskell)}*}
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   473
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   474
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   475
  \noindent In some cases it is desirable to have this
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   476
  pseudo-\qt{partiality} more explicitly, e.g.~as follows:
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   477
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   478
28564
haftmann
parents: 28562
diff changeset
   479
axiomatization %quote empty_queue :: 'a
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   480
28564
haftmann
parents: 28562
diff changeset
   481
function %quote strict_dequeue' :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   482
  "strict_dequeue' (Queue xs []) = (if xs = [] then empty_queue
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   483
     else strict_dequeue' (Queue [] (rev xs)))"
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   484
  | "strict_dequeue' (Queue xs (y # ys)) =
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   485
       (y, Queue xs ys)"
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   486
by pat_completeness auto
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   487
28564
haftmann
parents: 28562
diff changeset
   488
termination %quote strict_dequeue'
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   489
by (relation "measure (\<lambda>q::'a queue. case q of Queue xs ys \<Rightarrow> length xs)") simp_all
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   490
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   491
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   492
  \noindent For technical reasons the definition of
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   493
  @{const strict_dequeue'} is more involved since it requires
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   494
  a manual termination proof.  Apart from that observe that
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   495
  on the right hand side of its first equation the constant
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   496
  @{const empty_queue} occurs which is unspecified.
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   497
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
   498
  Normally, if constants without any code equations occur in
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   499
  a program, the code generator complains (since in most cases
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   500
  this is not what the user expects).  But such constants can also
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   501
  be thought of as function definitions with no equations which
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   502
  always fail, since there is never a successful pattern match
28593
f087237af65d continued codegen tutorial
haftmann
parents: 28564
diff changeset
   503
  on the left hand side.  In order to categorise a constant into
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   504
  that category explicitly, use @{command "code_abort"}:
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   505
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   506
28564
haftmann
parents: 28562
diff changeset
   507
code_abort %quote empty_queue
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   508
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   509
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   510
  \noindent Then the code generator will just insert an error or
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   511
  exception at the appropriate position:
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   512
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   513
28564
haftmann
parents: 28562
diff changeset
   514
text %quote {*@{code_stmts strict_dequeue' (consts) empty_queue strict_dequeue' (Haskell)}*}
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   515
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   516
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   517
  \noindent This feature however is rarely needed in practice.
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   518
  Note also that the @{text HOL} default setup already declares
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   519
  @{const undefined} as @{command "code_abort"}, which is most
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   520
  likely to be used in such situations.
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   521
*}
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   522
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   523
end
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   524