doc-src/Codegen/Thy/Further.thy
author haftmann
Mon, 20 Sep 2010 14:50:45 +0200
changeset 39559 e7d4923b9b1c
parent 39064 67f0cb151eb2
child 39664 0afaf89ab591
permissions -rw-r--r--
expand_fun_eq -> fun_eq_iff
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
     1
theory Further
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
     2
imports Setup
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
28447
haftmann
parents: 28419
diff changeset
     5
section {* Further issues \label{sec:further} *}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
     6
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
     7
subsection {* Modules namespace *}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
     8
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
     9
text {*
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    10
  When invoking the @{command export_code} command it is possible to leave
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    11
  out the @{keyword "module_name"} part;  then code is distributed over
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    12
  different modules, where the module name space roughly is induced
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    13
  by the Isabelle theory name space.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    14
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    15
  Then sometimes the awkward situation occurs that dependencies between
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    16
  definitions introduce cyclic dependencies between modules, which in the
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    17
  @{text Haskell} world leaves you to the mercy of the @{text Haskell} implementation
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    18
  you are using,  while for @{text SML}/@{text OCaml} code generation is not possible.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    19
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    20
  A solution is to declare module names explicitly.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    21
  Let use assume the three cyclically dependent
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    22
  modules are named \emph{A}, \emph{B} and \emph{C}.
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    23
  Then, by stating
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    24
*}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    25
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    26
code_modulename %quote SML
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    27
  A ABC
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    28
  B ABC
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    29
  C ABC
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    30
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    31
text {*\noindent
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    32
  we explicitly map all those modules on \emph{ABC},
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    33
  resulting in an ad-hoc merge of this three modules
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    34
  at serialisation time.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    35
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    36
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    37
subsection {* Locales and interpretation *}
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    38
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    39
text {*
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    40
  A technical issue comes to surface when generating code from
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    41
  specifications stemming from locale interpretation.
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    42
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    43
  Let us assume a locale specifying a power operation
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    44
  on arbitrary types:
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    45
*}
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    46
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    47
locale %quote power =
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    48
  fixes power :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    49
  assumes power_commute: "power x \<circ> power y = power y \<circ> power x"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    50
begin
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    51
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    52
text {*
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    53
  \noindent Inside that locale we can lift @{text power} to exponent lists
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    54
  by means of specification relative to that locale:
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    55
*}
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    56
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    57
primrec %quote powers :: "'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    58
  "powers [] = id"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    59
| "powers (x # xs) = power x \<circ> powers xs"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    60
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    61
lemma %quote powers_append:
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    62
  "powers (xs @ ys) = powers xs \<circ> powers ys"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    63
  by (induct xs) simp_all
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    64
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    65
lemma %quote powers_power:
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    66
  "powers xs \<circ> power x = power x \<circ> powers xs"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    67
  by (induct xs)
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    68
    (simp_all del: o_apply id_apply add: o_assoc [symmetric],
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    69
      simp del: o_apply add: o_assoc power_commute)
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    70
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    71
lemma %quote powers_rev:
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    72
  "powers (rev xs) = powers xs"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    73
    by (induct xs) (simp_all add: powers_append powers_power)
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    74
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    75
end %quote
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    76
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    77
text {*
38505
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38437
diff changeset
    78
  After an interpretation of this locale (say, @{command_def
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    79
  interpretation} @{text "fun_power:"} @{term [source] "power (\<lambda>n (f ::
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    80
  'a \<Rightarrow> 'a). f ^^ n)"}), one would expect to have a constant @{text
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    81
  "fun_power.powers :: nat list \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"} for which code
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    82
  can be generated.  But this not the case: internally, the term
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    83
  @{text "fun_power.powers"} is an abbreviation for the foundational
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    84
  term @{term [source] "power.powers (\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)"}
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    85
  (see \cite{isabelle-locale} for the details behind).
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    86
37836
2bcce92be291 adjusted; fixed typo
haftmann
parents: 37432
diff changeset
    87
  Fortunately, with minor effort the desired behaviour can be achieved.
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    88
  First, a dedicated definition of the constant on which the local @{text "powers"}
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    89
  after interpretation is supposed to be mapped on:
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    90
*}
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    91
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    92
definition %quote funpows :: "nat list \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    93
  [code del]: "funpows = power.powers (\<lambda>n f. f ^^ n)"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    94
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    95
text {*
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    96
  \noindent In general, the pattern is @{text "c = t"} where @{text c} is
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    97
  the name of the future constant and @{text t} the foundational term
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    98
  corresponding to the local constant after interpretation.
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
    99
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   100
  The interpretation itself is enriched with an equation @{text "t = c"}:
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   101
*}
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   102
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   103
interpretation %quote fun_power: power "\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n" where
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   104
  "power.powers (\<lambda>n f. f ^^ n) = funpows"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   105
  by unfold_locales
39559
e7d4923b9b1c expand_fun_eq -> fun_eq_iff
haftmann
parents: 39064
diff changeset
   106
    (simp_all add: fun_eq_iff funpow_mult mult_commute funpows_def)
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   107
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   108
text {*
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   109
  \noindent This additional equation is trivially proved by the definition
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   110
  itself.
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   111
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   112
  After this setup procedure, code generation can continue as usual:
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   113
*}
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   114
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   115
text %quote {*@{code_stmts funpows (consts) Nat.funpow funpows (Haskell)}*}
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   116
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   117
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   118
subsection {* Imperative data structures *}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   119
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   120
text {*
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   121
  If you consider imperative data structures as inevitable for a
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   122
  specific application, you should consider \emph{Imperative
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   123
  Functional Programming with Isabelle/HOL}
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   124
  \cite{bulwahn-et-al:2008:imperative}; the framework described there
39064
67f0cb151eb2 avoid theory Imperative_HOL altogether
haftmann
parents: 38510
diff changeset
   125
  is available in session @{text Imperative_HOL}.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   126
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   127
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   128
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   129
subsection {* ML system interfaces \label{sec:ml} *}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   130
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   131
text {*
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   132
  Since the code generator framework not only aims to provide a nice
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   133
  Isar interface but also to form a base for code-generation-based
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   134
  applications, here a short description of the most fundamental ML
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   135
  interfaces.
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   136
*}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   137
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   138
subsubsection {* Managing executable content *}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   139
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   140
text %mlref {*
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   141
  \begin{mldecls}
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   142
  @{index_ML Code.read_const: "theory -> string -> string"} \\
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   143
  @{index_ML Code.add_eqn: "thm -> theory -> theory"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   144
  @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   145
  @{index_ML Code_Preproc.map_pre: "(simpset -> simpset) -> theory -> theory"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   146
  @{index_ML Code_Preproc.map_post: "(simpset -> simpset) -> theory -> theory"} \\
38507
06728ef076a7 output whitespace tuning
haftmann
parents: 38505
diff changeset
   147
  @{index_ML Code_Preproc.add_functrans: "
06728ef076a7 output whitespace tuning
haftmann
parents: 38505
diff changeset
   148
    string * (theory -> (thm * bool) list -> (thm * bool) list option)
06728ef076a7 output whitespace tuning
haftmann
parents: 38505
diff changeset
   149
      -> theory -> theory"} \\
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   150
  @{index_ML Code_Preproc.del_functrans: "string -> theory -> theory"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   151
  @{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   152
  @{index_ML Code.get_type: "theory -> string
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   153
    -> (string * sort) list * ((string * string list) * typ list) list"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   154
  @{index_ML Code.get_type_of_constr_or_abstr: "theory -> string -> (string * bool) option"}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   155
  \end{mldecls}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   156
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   157
  \begin{description}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   158
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   159
  \item @{ML Code.read_const}~@{text thy}~@{text s}
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   160
     reads a constant as a concrete term expression @{text s}.
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   161
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   162
  \item @{ML Code.add_eqn}~@{text "thm"}~@{text "thy"} adds function
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   163
     theorem @{text "thm"} to executable content.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   164
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   165
  \item @{ML Code.del_eqn}~@{text "thm"}~@{text "thy"} removes function
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   166
     theorem @{text "thm"} from executable content, if present.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   167
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   168
  \item @{ML Code_Preproc.map_pre}~@{text "f"}~@{text "thy"} changes
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   169
     the preprocessor simpset.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   170
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   171
  \item @{ML Code_Preproc.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   172
     function transformer @{text f} (named @{text name}) to executable content;
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   173
     @{text f} is a transformer of the code equations belonging
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   174
     to a certain function definition, depending on the
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   175
     current theory context.  Returning @{text NONE} indicates that no
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   176
     transformation took place;  otherwise, the whole process will be iterated
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   177
     with the new code equations.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   178
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   179
  \item @{ML Code_Preproc.del_functrans}~@{text "name"}~@{text "thy"} removes
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   180
     function transformer named @{text name} from executable content.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   181
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   182
  \item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   183
     a datatype to executable content, with generation
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   184
     set @{text cs}.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   185
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   186
  \item @{ML Code.get_type_of_constr_or_abstr}~@{text "thy"}~@{text "const"}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   187
     returns type constructor corresponding to
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   188
     constructor @{text const}; returns @{text NONE}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   189
     if @{text const} is no constructor.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   190
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   191
  \end{description}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   192
*}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   193
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   194
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   195
subsubsection {* Data depending on the theory's executable content *}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   196
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   197
text {*
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   198
  Implementing code generator applications on top
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   199
  of the framework set out so far usually not only
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   200
  involves using those primitive interfaces
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   201
  but also storing code-dependent data and various
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   202
  other things.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   203
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   204
  Due to incrementality of code generation, changes in the
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   205
  theory's executable content have to be propagated in a
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   206
  certain fashion.  Additionally, such changes may occur
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   207
  not only during theory extension but also during theory
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   208
  merge, which is a little bit nasty from an implementation
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   209
  point of view.  The framework provides a solution
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   210
  to this technical challenge by providing a functorial
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   211
  data slot @{ML_functor Code_Data}; on instantiation
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   212
  of this functor, the following types and operations
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   213
  are required:
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   214
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   215
  \medskip
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   216
  \begin{tabular}{l}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   217
  @{text "type T"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   218
  @{text "val empty: T"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   219
  \end{tabular}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   220
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   221
  \begin{description}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   222
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   223
  \item @{text T} the type of data to store.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   224
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   225
  \item @{text empty} initial (empty) data.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   226
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   227
  \end{description}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   228
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   229
  \noindent An instance of @{ML_functor Code_Data} provides the following
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   230
  interface:
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   231
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   232
  \medskip
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   233
  \begin{tabular}{l}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   234
  @{text "change: theory \<rightarrow> (T \<rightarrow> T) \<rightarrow> T"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   235
  @{text "change_yield: theory \<rightarrow> (T \<rightarrow> 'a * T) \<rightarrow> 'a * T"}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   236
  \end{tabular}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   237
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   238
  \begin{description}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   239
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   240
  \item @{text change} update of current data (cached!)
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   241
    by giving a continuation.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   242
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   243
  \item @{text change_yield} update with side result.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   244
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   245
  \end{description}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   246
*}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   247
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   248
end