doc-src/Codegen/Thy/Program.thy
author haftmann
Mon, 14 Jun 2010 15:27:09 +0200
changeset 37427 e482f206821e
parent 37211 32a6f471f090
child 37612 48fed6598be9
permissions -rw-r--r--
added reference
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]:
32a6f471f090 clarified
haftmann
parents: 36259
diff changeset
   177
  "x \<in> set xs \<longleftrightarrow> member xs x" by (fact in_set_code)
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]:
32a6f471f090 clarified
haftmann
parents: 36259
diff changeset
   191
  "xs = [] \<longleftrightarrow> List.null xs" by (fact empty_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 {*
28428
haftmann
parents: 28419
diff changeset
   342
  \noindent The membership test during preprocessing 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
33926
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   432
subsection {* Inductive Predicates *}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   433
(*<*)
36176
3fe7e97ccca8 replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents: 34155
diff changeset
   434
hide_const append
33926
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   435
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   436
inductive append
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   437
where
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   438
  "append [] ys ys"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   439
| "append xs ys zs ==> append (x # xs) ys (x # zs)"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   440
(*>*)
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   441
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   442
To execute inductive predicates, a special preprocessor, the predicate
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   443
 compiler, generates code equations from the introduction rules of the predicates.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   444
The mechanisms of this compiler are described in \cite{Berghofer-Bulwahn-Haftmann:2009:TPHOL}.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   445
Consider the simple predicate @{const append} given by these two
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   446
introduction rules:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   447
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   448
text %quote {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   449
@{thm append.intros(1)[of ys]}\\
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   450
\noindent@{thm append.intros(2)[of xs ys zs x]}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   451
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   452
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   453
\noindent To invoke the compiler, simply use @{command_def "code_pred"}:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   454
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   455
code_pred %quote append .
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   456
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   457
\noindent The @{command "code_pred"} command takes the name
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   458
of the inductive predicate and then you put a period to discharge
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   459
a trivial correctness proof. 
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   460
The compiler infers possible modes 
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   461
for the predicate and produces the derived code equations.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   462
Modes annotate which (parts of the) arguments are to be taken as input,
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   463
and which output. Modes are similar to types, but use the notation @{text "i"}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   464
for input and @{text "o"} for output.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   465
 
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   466
For @{term "append"}, the compiler can infer the following modes:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   467
\begin{itemize}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   468
\item @{text "i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool"}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   469
\item @{text "i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool"}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   470
\item @{text "o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool"}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   471
\end{itemize}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   472
You can compute sets of predicates using @{command_def "values"}:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   473
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   474
values %quote "{zs. append [(1::nat),2,3] [4,5] zs}"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   475
text {* \noindent outputs @{text "{[1, 2, 3, 4, 5]}"}, and *}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   476
values %quote "{(xs, ys). append xs ys [(2::nat),3]}"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   477
text {* \noindent outputs @{text "{([], [2, 3]), ([2], [3]), ([2, 3], [])}"}. *}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   478
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   479
\noindent If you are only interested in the first elements of the set
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   480
comprehension (with respect to a depth-first search on the introduction rules), you can
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   481
pass an argument to
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   482
@{command "values"} to specify the number of elements you want:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   483
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   484
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   485
values %quote 1 "{(xs, ys). append xs ys [(1::nat),2,3,4]}"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   486
values %quote 3 "{(xs, ys). append xs ys [(1::nat),2,3,4]}"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   487
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   488
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   489
\noindent The @{command "values"} command can only compute set
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   490
 comprehensions for which a mode has been inferred.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   491
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   492
The code equations for a predicate are made available as theorems with
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   493
 the suffix @{text "equation"}, and can be inspected with:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   494
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   495
thm %quote append.equation
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   496
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   497
\noindent More advanced options are described in the following subsections.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   498
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   499
subsubsection {* Alternative names for functions *}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   500
text {* 
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   501
By default, the functions generated from a predicate are named after the predicate with the
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   502
mode mangled into the name (e.g., @{text "append_i_i_o"}).
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   503
You can specify your own names as follows:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   504
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   505
code_pred %quote (modes: i => i => o => bool as concat,
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   506
  o => o => i => bool as split,
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   507
  i => o => i => bool as suffix) append .
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   508
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   509
subsubsection {* Alternative introduction rules *}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   510
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   511
Sometimes the introduction rules of an predicate are not executable because they contain
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   512
non-executable constants or specific modes could not be inferred.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   513
It is also possible that the introduction rules yield a function that loops forever
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   514
due to the execution in a depth-first search manner.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   515
Therefore, you can declare alternative introduction rules for predicates with the attribute
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   516
@{attribute "code_pred_intro"}.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   517
For example, the transitive closure is defined by: 
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   518
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   519
text %quote {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   520
@{thm tranclp.intros(1)[of r a b]}\\
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   521
\noindent @{thm tranclp.intros(2)[of r a b c]}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   522
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   523
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   524
\noindent These rules do not suit well for executing the transitive closure 
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   525
with the mode @{text "(i \<Rightarrow> o \<Rightarrow> bool) \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool"}, as the second rule will
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   526
cause an infinite loop in the recursive call.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   527
This can be avoided using the following alternative rules which are
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   528
declared to the predicate compiler by the attribute @{attribute "code_pred_intro"}:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   529
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   530
lemma %quote [code_pred_intro]:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   531
  "r a b \<Longrightarrow> r\<^sup>+\<^sup>+ a b"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   532
  "r a b \<Longrightarrow> r\<^sup>+\<^sup>+ b c \<Longrightarrow> r\<^sup>+\<^sup>+ a c"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   533
by auto
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   534
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   535
\noindent After declaring all alternative rules for the transitive closure,
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   536
you invoke @{command "code_pred"} as usual.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   537
As you have declared alternative rules for the predicate, you are urged to prove that these
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   538
introduction rules are complete, i.e., that you can derive an elimination rule for the
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   539
alternative rules:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   540
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   541
code_pred %quote tranclp
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   542
proof -
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   543
  case tranclp
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   544
  from this converse_tranclpE[OF this(1)] show thesis by metis
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   545
qed
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   546
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   547
\noindent Alternative rules can also be used for constants that have not
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   548
been defined inductively. For example, the lexicographic order which
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   549
is defined as: *}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   550
text %quote {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   551
@{thm [display] lexord_def[of "r"]}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   552
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   553
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   554
\noindent To make it executable, you can derive the following two rules and prove the
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   555
elimination rule:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   556
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   557
(*<*)
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   558
lemma append: "append xs ys zs = (xs @ ys = zs)"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   559
by (induct xs arbitrary: ys zs) (auto elim: append.cases intro: append.intros)
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   560
(*>*)
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   561
lemma %quote [code_pred_intro]:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   562
  "append xs (a # v) ys \<Longrightarrow> lexord r (xs, ys)"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   563
(*<*)unfolding lexord_def Collect_def by (auto simp add: append)(*>*)
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   564
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   565
lemma %quote [code_pred_intro]:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   566
  "append u (a # v) xs \<Longrightarrow> append u (b # w) ys \<Longrightarrow> r (a, b)
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   567
  \<Longrightarrow> lexord r (xs, ys)"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   568
(*<*)unfolding lexord_def Collect_def append mem_def apply simp
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   569
apply (rule disjI2) by auto
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   570
(*>*)
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   571
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   572
code_pred %quote lexord
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   573
(*<*)
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   574
proof -
36259
9f9b9b14cc7a adopting documentation of the predicate compiler
bulwahn
parents: 36176
diff changeset
   575
  fix r xs ys
9f9b9b14cc7a adopting documentation of the predicate compiler
bulwahn
parents: 36176
diff changeset
   576
  assume lexord: "lexord r (xs, ys)"
9f9b9b14cc7a adopting documentation of the predicate compiler
bulwahn
parents: 36176
diff changeset
   577
  assume 1: "\<And> r' xs' ys' a v. r = r' \<Longrightarrow> (xs, ys) = (xs', ys') \<Longrightarrow> append xs' (a # v) ys' \<Longrightarrow> thesis"
9f9b9b14cc7a adopting documentation of the predicate compiler
bulwahn
parents: 36176
diff changeset
   578
  assume 2: "\<And> r' xs' ys' u a v b w. r = r' \<Longrightarrow> (xs, ys) = (xs', ys') \<Longrightarrow> append u (a # v) xs' \<Longrightarrow> append u (b # w) ys' \<Longrightarrow> r' (a, b) \<Longrightarrow> thesis"
33926
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   579
  {
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   580
    assume "\<exists>a v. ys = xs @ a # v"
36259
9f9b9b14cc7a adopting documentation of the predicate compiler
bulwahn
parents: 36176
diff changeset
   581
    from this 1 have thesis
33926
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   582
        by (fastsimp simp add: append)
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   583
  } moreover
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   584
  {
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   585
    assume "\<exists>u a b v w. (a, b) \<in> r \<and> xs = u @ a # v \<and> ys = u @ b # w"
36259
9f9b9b14cc7a adopting documentation of the predicate compiler
bulwahn
parents: 36176
diff changeset
   586
    from this 2 have thesis by (fastsimp simp add: append mem_def)
33926
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   587
  } moreover
36259
9f9b9b14cc7a adopting documentation of the predicate compiler
bulwahn
parents: 36176
diff changeset
   588
  note lexord
33926
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   589
  ultimately show thesis
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   590
    unfolding lexord_def
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   591
    by (fastsimp simp add: Collect_def)
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   592
qed
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   593
(*>*)
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   594
subsubsection {* Options for values *}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   595
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   596
In the presence of higher-order predicates, multiple modes for some predicate could be inferred
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   597
that are not disambiguated by the pattern of the set comprehension.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   598
To disambiguate the modes for the arguments of a predicate, you can state
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   599
the modes explicitly in the @{command "values"} command. 
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   600
Consider the simple predicate @{term "succ"}:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   601
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   602
inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   603
where
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   604
  "succ 0 (Suc 0)"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   605
| "succ x y \<Longrightarrow> succ (Suc x) (Suc y)"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   606
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   607
code_pred succ .
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   608
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   609
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   610
\noindent For this, the predicate compiler can infer modes @{text "o \<Rightarrow> o \<Rightarrow> bool"}, @{text "i \<Rightarrow> o \<Rightarrow> bool"},
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   611
  @{text "o \<Rightarrow> i \<Rightarrow> bool"} and @{text "i \<Rightarrow> i \<Rightarrow> bool"}.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   612
The invocation of @{command "values"} @{text "{n. tranclp succ 10 n}"} loops, as multiple
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   613
modes for the predicate @{text "succ"} are possible and here the first mode @{text "o \<Rightarrow> o \<Rightarrow> bool"}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   614
is chosen. To choose another mode for the argument, you can declare the mode for the argument between
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   615
the @{command "values"} and the number of elements.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   616
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   617
values %quote [mode: i => o => bool] 20 "{n. tranclp succ 10 n}"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   618
values %quote [mode: o => i => bool] 10 "{n. tranclp succ n 10}"
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   619
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   620
subsubsection {* Embedding into functional code within Isabelle/HOL *}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   621
text {*
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   622
To embed the computation of an inductive predicate into functions that are defined in Isabelle/HOL,
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   623
you have a number of options:
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   624
\begin{itemize}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   625
\item You want to use the first-order predicate with the mode
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   626
  where all arguments are input. Then you can use the predicate directly, e.g.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   627
\begin{quote}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   628
  @{text "valid_suffix ys zs = "}\\
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   629
  @{text "(if append [Suc 0, 2] ys zs then Some ys else None)"}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   630
\end{quote}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   631
\item If you know that the execution returns only one value (it is deterministic), then you can
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   632
  use the combinator @{term "Predicate.the"}, e.g., a functional concatenation of lists
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   633
  is defined with
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   634
\begin{quote}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   635
  @{term "functional_concat xs ys = Predicate.the (append_i_i_o xs ys)"}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   636
\end{quote}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   637
  Note that if the evaluation does not return a unique value, it raises a run-time error
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   638
  @{term "not_unique"}.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   639
\end{itemize}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   640
*}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   641
subsubsection {* Further Examples *}
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   642
text {* Further examples for compiling inductive predicates can be found in
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   643
the @{text "HOL/ex/Predicate_Compile_ex"} theory file.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   644
There are also some examples in the Archive of Formal Proofs, notably
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   645
in the @{text "POPLmark-deBruijn"} and the @{text "FeatherweightJava"} sessions.
dd017d9db05f adding subsection about the predicate compiler to the code generator tutorial
bulwahn
parents: 33707
diff changeset
   646
*}
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   647
end