doc-src/Codegen/Thy/Program.thy
author wenzelm
Tue, 03 Aug 2010 22:28:43 +0200
changeset 38142 c202426474c3
parent 37612 48fed6598be9
permissions -rw-r--r--
more precise CRITICAL sections; tuned;
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
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
    11
  @{command definition}, @{command primrec} and @{command fun}
28447
haftmann
parents: 28428
diff changeset
    12
  statements are used for code generation.  This default behaviour
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
    13
  can be changed, e.g.\ by providing different code equations.
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
    14
  The customisations shown in this section are \emph{safe}
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
    15
  as regards correctness: all programs that can be generated are partially
28447
haftmann
parents: 28428
diff changeset
    16
  correct.
haftmann
parents: 28428
diff changeset
    17
*}
haftmann
parents: 28428
diff changeset
    18
haftmann
parents: 28428
diff changeset
    19
subsection {* Selecting code equations *}
haftmann
parents: 28428
diff changeset
    20
haftmann
parents: 28428
diff changeset
    21
text {*
haftmann
parents: 28428
diff changeset
    22
  Coming back to our introductory example, we
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
    23
  could provide an alternative code equations for @{const dequeue}
28447
haftmann
parents: 28428
diff changeset
    24
  explicitly:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    25
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    26
28564
haftmann
parents: 28562
diff changeset
    27
lemma %quote [code]:
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
    28
  "dequeue (AQueue xs []) =
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
    29
     (if xs = [] then (None, AQueue [] [])
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
    30
       else dequeue (AQueue [] (rev xs)))"
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
    31
  "dequeue (AQueue xs (y # ys)) =
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
    32
     (Some y, AQueue xs ys)"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    33
  by (cases xs, simp_all) (cases "rev xs", simp_all)
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    34
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    35
text {*
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28462
diff changeset
    36
  \noindent The annotation @{text "[code]"} is an @{text Isar}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    37
  @{text attribute} which states that the given theorems should be
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
    38
  considered as code equations for a @{text fun} statement --
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    39
  the corresponding constant is determined syntactically.  The resulting code:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    40
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    41
28564
haftmann
parents: 28562
diff changeset
    42
text %quote {*@{code_stmts dequeue (consts) dequeue (Haskell)}*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    43
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    44
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    45
  \noindent You may note that the equality test @{term "xs = []"} has been
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    46
  replaced by the predicate @{term "null xs"}.  This is due to the default
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    47
  setup in the \qn{preprocessor} to be discussed further below (\secref{sec:preproc}).
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    48
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    49
  Changing the default constructor set of datatypes is also
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
    50
  possible.  See \secref{sec:datatypes} for an example.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    51
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    52
  As told in \secref{sec:concept}, code generation is based
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    53
  on a structured collection of code theorems.
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
    54
  This collection
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    55
  may be inspected using the @{command code_thms} command:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    56
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    57
28564
haftmann
parents: 28562
diff changeset
    58
code_thms %quote dequeue
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    59
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    60
text {*
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
    61
  \noindent prints a table with \emph{all} code equations
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    62
  for @{const dequeue}, including
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
    63
  \emph{all} code equations those equations depend
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    64
  on recursively.
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    65
  
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    66
  Similarly, the @{command code_deps} command shows a graph
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
    67
  visualising dependencies between code equations.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    68
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    69
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    70
subsection {* @{text class} and @{text instantiation} *}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    71
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    72
text {*
28447
haftmann
parents: 28428
diff changeset
    73
  Concerning type classes and code generation, let us examine an example
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    74
  from abstract algebra:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    75
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    76
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
    77
class %quote semigroup =
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    78
  fixes mult :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<otimes>" 70)
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    79
  assumes assoc: "(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    80
28564
haftmann
parents: 28562
diff changeset
    81
class %quote monoid = semigroup +
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    82
  fixes neutral :: 'a ("\<one>")
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    83
  assumes neutl: "\<one> \<otimes> x = x"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    84
    and neutr: "x \<otimes> \<one> = x"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    85
28564
haftmann
parents: 28562
diff changeset
    86
instantiation %quote nat :: monoid
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    87
begin
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    88
28564
haftmann
parents: 28562
diff changeset
    89
primrec %quote mult_nat where
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    90
    "0 \<otimes> n = (0\<Colon>nat)"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    91
  | "Suc m \<otimes> n = n + m \<otimes> n"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    92
28564
haftmann
parents: 28562
diff changeset
    93
definition %quote neutral_nat where
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    94
  "\<one> = Suc 0"
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
    95
28564
haftmann
parents: 28562
diff changeset
    96
lemma %quote add_mult_distrib:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    97
  fixes n m q :: nat
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    98
  shows "(n + m) \<otimes> q = n \<otimes> q + m \<otimes> q"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    99
  by (induct n) simp_all
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   100
28564
haftmann
parents: 28562
diff changeset
   101
instance %quote proof
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   102
  fix m n q :: nat
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   103
  show "m \<otimes> n \<otimes> q = m \<otimes> (n \<otimes> q)"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   104
    by (induct m) (simp_all add: add_mult_distrib)
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   105
  show "\<one> \<otimes> n = n"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   106
    by (simp add: neutral_nat_def)
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   107
  show "m \<otimes> \<one> = m"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   108
    by (induct m) (simp_all add: neutral_nat_def)
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   109
qed
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   110
28564
haftmann
parents: 28562
diff changeset
   111
end %quote
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   112
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   113
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   114
  \noindent We define the natural operation of the natural numbers
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   115
  on monoids:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   116
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   117
28564
haftmann
parents: 28562
diff changeset
   118
primrec %quote (in monoid) pow :: "nat \<Rightarrow> 'a \<Rightarrow> 'a" where
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   119
    "pow 0 a = \<one>"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   120
  | "pow (Suc n) a = a \<otimes> pow n a"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   121
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   122
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   123
  \noindent This we use to define the discrete exponentiation function:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   124
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   125
28564
haftmann
parents: 28562
diff changeset
   126
definition %quote bexp :: "nat \<Rightarrow> nat" where
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   127
  "bexp n = pow n (Suc (Suc 0))"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   128
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   129
text {*
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   130
  \noindent The corresponding code in Haskell uses that language's native classes:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   131
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   132
28564
haftmann
parents: 28562
diff changeset
   133
text %quote {*@{code_stmts bexp (Haskell)}*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   134
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   135
text {*
28447
haftmann
parents: 28428
diff changeset
   136
  \noindent This is a convenient place to show how explicit dictionary construction
37427
e482f206821e added reference
haftmann
parents: 37211
diff changeset
   137
  manifests in generated code (here, the same example in @{text SML})
e482f206821e added reference
haftmann
parents: 37211
diff changeset
   138
  \cite{Haftmann-Nipkow:2010:code}:
28419
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 {*
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   144
  \noindent Note the parameters with trailing underscore (@{verbatim "A_"}),
28447
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
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   156
  The \emph{simpset} can apply the full generality of the
32000
6f07563dc8a9 updated to changes in sources; tuned
haftmann
parents: 31254
diff changeset
   157
  Isabelle simplifier.  Due to the interpretation of theorems as code
6f07563dc8a9 updated to changes in sources; tuned
haftmann
parents: 31254
diff changeset
   158
  equations, rewrites are applied to the right hand side and the
6f07563dc8a9 updated to changes in sources; tuned
haftmann
parents: 31254
diff changeset
   159
  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
   160
  constant heading the left hand side.  An important special case are
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   161
  \emph{unfold theorems},  which may be declared and removed using
32000
6f07563dc8a9 updated to changes in sources; tuned
haftmann
parents: 31254
diff changeset
   162
  the @{attribute code_unfold} or \emph{@{attribute code_unfold} del}
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   163
  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
37211
32a6f471f090 clarified
haftmann
parents: 36259
diff changeset
   176
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
   177
  "x \<in> set xs \<longleftrightarrow> List.member xs x" by (fact in_set_member)
28419
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
37211
32a6f471f090 clarified
haftmann
parents: 36259
diff changeset
   183
lemma %quote [code_unfold]:
32a6f471f090 clarified
haftmann
parents: 36259
diff changeset
   184
  "1 = Suc 0" by (fact One_nat_def)
28419
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
37211
32a6f471f090 clarified
haftmann
parents: 36259
diff changeset
   190
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
   191
  "xs = [] \<longleftrightarrow> List.null xs" by (fact eq_Nil_null)
28419
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
31254
03a35fbc9dc6 documented print_codeproc command
haftmann
parents: 30938
diff changeset
   207
  the @{command print_codeproc} command.
28419
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}
32000
6f07563dc8a9 updated to changes in sources; tuned
haftmann
parents: 31254
diff changeset
   213
6f07563dc8a9 updated to changes in sources; tuned
haftmann
parents: 31254
diff changeset
   214
    Attribute @{attribute code_unfold} also applies to the
6f07563dc8a9 updated to changes in sources; tuned
haftmann
parents: 31254
diff changeset
   215
    preprocessor of the ancient @{text "SML code generator"}; in case
6f07563dc8a9 updated to changes in sources; tuned
haftmann
parents: 31254
diff changeset
   216
    this is not what you intend, use @{attribute code_inline} instead.
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
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   235
  theorems about it.  Therefore, here is a simple, straightforward
29794
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
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   261
text {*
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   262
  \noindent Here we define a \qt{constructor} @{const "AQueue"} which
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   263
  is defined in terms of @{text "Queue"} and interprets its arguments
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   264
  according to what the \emph{content} of an amortised queue is supposed
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   265
  to be.  Equipped with this, we are able to prove the following equations
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   266
  for our primitive queue operations which \qt{implement} the simple
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   267
  queues in an amortised fashion:
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   268
*}
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   269
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   270
lemma %quote empty_AQueue [code]:
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   271
  "empty = AQueue [] []"
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   272
  unfolding AQueue_def empty_def by simp
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   273
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   274
lemma %quote enqueue_AQueue [code]:
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   275
  "enqueue x (AQueue xs ys) = AQueue (x # xs) ys"
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   276
  unfolding AQueue_def by simp
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   277
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   278
lemma %quote dequeue_AQueue [code]:
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   279
  "dequeue (AQueue xs []) =
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   280
    (if xs = [] then (None, AQueue [] [])
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   281
    else dequeue (AQueue [] (rev xs)))"
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   282
  "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)"
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   283
  unfolding AQueue_def by simp_all
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   284
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   285
text {*
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   286
  \noindent For completeness, we provide a substitute for the
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   287
  @{text case} combinator on queues:
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   288
*}
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   289
30227
853abb4853cc tuned manuals
haftmann
parents: 30226
diff changeset
   290
lemma %quote queue_case_AQueue [code]:
853abb4853cc tuned manuals
haftmann
parents: 30226
diff changeset
   291
  "queue_case f (AQueue xs ys) = f (ys @ rev xs)"
853abb4853cc tuned manuals
haftmann
parents: 30226
diff changeset
   292
  unfolding AQueue_def by simp
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   293
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   294
text {*
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   295
  \noindent The resulting code looks as expected:
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   296
*}
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   297
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   298
text %quote {*@{code_stmts empty enqueue dequeue (SML)}*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   299
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   300
text {*
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   301
  \noindent From this example, it can be glimpsed that using own
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   302
  constructor sets is a little delicate since it changes the set of
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   303
  valid patterns for values of that type.  Without going into much
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   304
  detail, here some practical hints:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   305
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   306
  \begin{itemize}
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   307
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   308
    \item When changing the constructor set for datatypes, take care
30227
853abb4853cc tuned manuals
haftmann
parents: 30226
diff changeset
   309
      to provide alternative equations for the @{text case} combinator.
29794
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   310
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   311
    \item Values in the target language need not to be normalised --
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   312
      different values in the target language may represent the same
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   313
      value in the logic.
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   314
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   315
    \item Usually, a good methodology to deal with the subtleties of
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   316
      pattern matching is to see the type as an abstract type: provide
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   317
      a set of operations which operate on the concrete representation
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   318
      of the type, and derive further operations by combinations of
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   319
      these primitive ones, without relying on a particular
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   320
      representation.
32d00a2a6f28 added stub about datatype abstraction
haftmann
parents: 29560
diff changeset
   321
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   322
  \end{itemize}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   323
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   324
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   325
30938
c6c9359e474c wellsortedness is no issue for a user manual any more
haftmann
parents: 30227
diff changeset
   326
subsection {* Equality *}
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   327
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   328
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   329
  Surely you have already noticed how equality is treated
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   330
  by the code generator:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   331
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   332
28564
haftmann
parents: 28562
diff changeset
   333
primrec %quote collect_duplicates :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
28447
haftmann
parents: 28428
diff changeset
   334
  "collect_duplicates xs ys [] = xs"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   335
  | "collect_duplicates xs ys (z#zs) = (if z \<in> set xs
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   336
      then if z \<in> set ys
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   337
        then collect_duplicates xs ys zs
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   338
        else collect_duplicates xs (z#ys) zs
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   339
      else collect_duplicates (z#xs) (z#ys) zs)"
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   340
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   341
text {*
37612
48fed6598be9 adapted to reorganization of auxiliary list operations; split off predicate compiler into separate theory
haftmann
parents: 37427
diff changeset
   342
  \noindent During preprocessing, the membership test is rewritten,
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   343
  resulting in @{const List.member}, which itself
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   344
  performs an explicit equality check.
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   345
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   346
28564
haftmann
parents: 28562
diff changeset
   347
text %quote {*@{code_stmts collect_duplicates (SML)}*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   348
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   349
text {*
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   350
  \noindent Obviously, polymorphic equality is implemented the Haskell
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   351
  way using a type class.  How is this achieved?  HOL introduces
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   352
  an explicit class @{class eq} with a corresponding operation
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   353
  @{const eq_class.eq} such that @{thm eq [no_vars]}.
28447
haftmann
parents: 28428
diff changeset
   354
  The preprocessing framework does the rest by propagating the
29560
fa6c5d62adf5 "code equation" replaces "defining equation"
haftmann
parents: 28593
diff changeset
   355
  @{class eq} constraints through all dependent code equations.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   356
  For datatypes, instances of @{class eq} are implicitly derived
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   357
  when possible.  For other types, you may instantiate @{text eq}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   358
  manually like any other type class.
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   359
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   360
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   361
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   362
subsection {* Explicit partiality *}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   363
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   364
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   365
  Partiality usually enters the game by partial patterns, as
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   366
  in the following example, again for amortised queues:
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   367
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   368
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   369
definition %quote strict_dequeue :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   370
  "strict_dequeue q = (case dequeue q
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   371
    of (Some x, q') \<Rightarrow> (x, q'))"
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   372
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   373
lemma %quote strict_dequeue_AQueue [code]:
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   374
  "strict_dequeue (AQueue xs (y # ys)) = (y, AQueue xs ys)"
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   375
  "strict_dequeue (AQueue xs []) =
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   376
    (case rev xs of y # ys \<Rightarrow> (y, AQueue [] ys))"
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   377
  by (simp_all add: strict_dequeue_def dequeue_AQueue split: list.splits)
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   378
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   379
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   380
  \noindent In the corresponding code, there is no equation
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   381
  for the pattern @{term "AQueue [] []"}:
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   382
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   383
28564
haftmann
parents: 28562
diff changeset
   384
text %quote {*@{code_stmts strict_dequeue (consts) strict_dequeue (Haskell)}*}
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   385
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   386
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   387
  \noindent In some cases it is desirable to have this
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   388
  pseudo-\qt{partiality} more explicitly, e.g.~as follows:
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   389
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   390
28564
haftmann
parents: 28562
diff changeset
   391
axiomatization %quote empty_queue :: 'a
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   392
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   393
definition %quote strict_dequeue' :: "'a queue \<Rightarrow> 'a \<times> 'a queue" where
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   394
  "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
   395
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   396
lemma %quote strict_dequeue'_AQueue [code]:
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   397
  "strict_dequeue' (AQueue xs []) = (if xs = [] then empty_queue
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   398
     else strict_dequeue' (AQueue [] (rev xs)))"
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   399
  "strict_dequeue' (AQueue xs (y # ys)) =
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   400
     (y, AQueue xs ys)"
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   401
  by (simp_all add: strict_dequeue'_def dequeue_AQueue split: list.splits)
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   402
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   403
text {*
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   404
  Observe that on the right hand side of the definition of @{const
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   405
  "strict_dequeue'"}, the unspecified constant @{const empty_queue} occurs.
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   406
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   407
  Normally, if constants without any code equations occur in a
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   408
  program, the code generator complains (since in most cases this is
34155
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   409
  indeed an error).  But such constants can also be thought
14aaccb399b3 Polishing up the English
paulson
parents: 33926
diff changeset
   410
  of as function definitions which always fail,
29798
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   411
  since there is never a successful pattern match on the left hand
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   412
  side.  In order to categorise a constant into that category
6df726203e39 proper datatype abstraction example
haftmann
parents: 29796
diff changeset
   413
  explicitly, use @{command "code_abort"}:
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   414
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   415
28564
haftmann
parents: 28562
diff changeset
   416
code_abort %quote empty_queue
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   417
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   418
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   419
  \noindent Then the code generator will just insert an error or
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   420
  exception at the appropriate position:
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   421
*}
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   422
28564
haftmann
parents: 28562
diff changeset
   423
text %quote {*@{code_stmts strict_dequeue' (consts) empty_queue strict_dequeue' (Haskell)}*}
28462
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   424
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   425
text {*
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   426
  \noindent This feature however is rarely needed in practice.
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   427
  Note also that the @{text HOL} default setup already declares
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   428
  @{const undefined} as @{command "code_abort"}, which is most
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   429
  likely to be used in such situations.
6ec603695aaf added partiality section
haftmann
parents: 28456
diff changeset
   430
*}
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   431
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   432
end