src/Doc/Codegen/Further.thy
author wenzelm
Fri, 23 Jun 2017 14:38:32 +0200
changeset 66176 b51a40281016
parent 65981 e2c25346b156
child 66251 cd935b7cb3fb
permissions -rw-r--r--
tuned signature;
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
59378
065f349852e6 separate image for prerequisites of codegen tutorial
haftmann
parents: 59377
diff changeset
     2
imports Setup
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
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
     5
section \<open>Further issues \label{sec:further}\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
     6
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
     7
subsection \<open>Incorporating generated code directly into the system runtime -- @{text code_reflect}\<close>
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
     8
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
     9
subsubsection \<open>Static embedding of generated code into the system runtime\<close>
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    10
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    11
text \<open>
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    12
  The @{ML_antiquotation code} antiquotation is lightweight, but the generated code
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    13
  is only accessible while the ML section is processed.  Sometimes this
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    14
  is not appropriate, especially if the generated code contains datatype
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    15
  declarations which are shared with other parts of the system.  In these
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    16
  cases, @{command_def code_reflect} can be used:
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    17
\<close>
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    18
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    19
code_reflect %quote Sum_Type
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    20
  datatypes sum = Inl | Inr
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    21
  functions "Sum_Type.sum.projl" "Sum_Type.sum.projr"
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    22
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    23
text \<open>
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    24
  \noindent @{command code_reflect} takes a structure name and
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    25
  references to datatypes and functions; for these code is compiled
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    26
  into the named ML structure and the \emph{Eval} target is modified
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    27
  in a way that future code generation will reference these
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    28
  precompiled versions of the given datatypes and functions.  This
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    29
  also allows to refer to the referenced datatypes and functions from
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    30
  arbitrary ML code as well.
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    31
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    32
  A typical example for @{command code_reflect} can be found in the
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    33
  @{theory Predicate} theory.
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    34
\<close>
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    35
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    36
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    37
subsubsection \<open>Separate compilation\<close>
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    38
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    39
text \<open>
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    40
  For technical reasons it is sometimes necessary to separate
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    41
  generation and compilation of code which is supposed to be used in
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    42
  the system runtime.  For this @{command code_reflect} with an
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    43
  optional \<^theory_text>\<open>file\<close> argument can be used:
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    44
\<close>
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    45
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    46
code_reflect %quote Rat
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    47
  datatypes rat
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    48
  functions Fract
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    49
    "(plus :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(minus :: rat \<Rightarrow> rat \<Rightarrow> rat)"
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    50
    "(times :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(divide :: rat \<Rightarrow> rat \<Rightarrow> rat)"
65981
e2c25346b156 more robust -- avoid race condition wrt. Haskell output in $ISABELLE_TMP/examples/
wenzelm
parents: 65041
diff changeset
    51
  file "$ISABELLE_TMP/rat.ML"
65041
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    52
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    53
text \<open>
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    54
  \noindent This merely generates the referenced code to the given
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    55
  file which can be included into the system runtime later on.
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    56
\<close>
2525e680f94f basic documentation for computations
haftmann
parents: 63680
diff changeset
    57
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
    58
subsection \<open>Specialities of the @{text Scala} target language \label{sec:scala}\<close>
42096
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    59
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
    60
text \<open>
42096
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    61
  @{text Scala} deviates from languages of the ML family in a couple
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    62
  of aspects; those which affect code generation mainly have to do with
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    63
  @{text Scala}'s type system:
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    64
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    65
  \begin{itemize}
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    66
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    67
    \item @{text Scala} prefers tupled syntax over curried syntax.
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    68
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    69
    \item @{text Scala} sacrifices Hindely-Milner type inference for a
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    70
      much more rich type system with subtyping etc.  For this reason
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    71
      type arguments sometimes have to be given explicitly in square
46520
a0abc2ea815e corrected spelling
haftmann
parents: 46514
diff changeset
    72
      brackets (mimicking System F syntax).
42096
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    73
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    74
    \item In contrast to @{text Haskell} where most specialities of
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    75
      the type system are implemented using \emph{type classes},
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    76
      @{text Scala} provides a sophisticated system of \emph{implicit
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    77
      arguments}.
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    78
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    79
  \end{itemize}
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    80
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    81
  \noindent Concerning currying, the @{text Scala} serializer counts
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    82
  arguments in code equations to determine how many arguments
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    83
  shall be tupled; remaining arguments and abstractions in terms
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    84
  rather than function definitions are always curried.
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    85
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    86
  The second aspect affects user-defined adaptations with @{command
55372
3662c44d018c dropped legacy finally
haftmann
parents: 52665
diff changeset
    87
  code_printing}.  For regular terms, the @{text Scala} serializer prints
42096
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    88
  all type arguments explicitly.  For user-defined term adaptations
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    89
  this is only possible for adaptations which take no arguments: here
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    90
  the type arguments are just appended.  Otherwise they are ignored;
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    91
  hence user-defined adaptations for polymorphic constants have to be
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    92
  designed very carefully to avoid ambiguity.
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
    93
\<close>
42096
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    94
9f6652122963 added subsection on Scala specialities
haftmann
parents: 40752
diff changeset
    95
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
    96
subsection \<open>Modules namespace\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
    97
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
    98
text \<open>
40752
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
    99
  When invoking the @{command export_code} command it is possible to
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   100
  leave out the @{keyword "module_name"} part; then code is
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   101
  distributed over different modules, where the module name space
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   102
  roughly is induced by the Isabelle theory name space.
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   103
40752
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   104
  Then sometimes the awkward situation occurs that dependencies
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   105
  between definitions introduce cyclic dependencies between modules,
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   106
  which in the @{text Haskell} world leaves you to the mercy of the
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   107
  @{text Haskell} implementation you are using, while for @{text
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   108
  SML}/@{text OCaml} code generation is not possible.
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   109
40752
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   110
  A solution is to declare module names explicitly.  Let use assume
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   111
  the three cyclically dependent modules are named \emph{A}, \emph{B}
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   112
  and \emph{C}.  Then, by stating
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   113
\<close>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   114
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51717
diff changeset
   115
code_identifier %quote
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51717
diff changeset
   116
  code_module A \<rightharpoonup> (SML) ABC
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51717
diff changeset
   117
| code_module B \<rightharpoonup> (SML) ABC
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51717
diff changeset
   118
| code_module C \<rightharpoonup> (SML) ABC
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   119
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   120
text \<open>
40752
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   121
  \noindent we explicitly map all those modules on \emph{ABC},
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   122
  resulting in an ad-hoc merge of this three modules at serialisation
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   123
  time.
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   124
\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   125
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   126
subsection \<open>Locales and interpretation\<close>
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   127
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   128
text \<open>
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   129
  A technical issue comes to surface when generating code from
61891
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61779
diff changeset
   130
  specifications stemming from locale interpretation into global
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61779
diff changeset
   131
  theories.
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   132
40752
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   133
  Let us assume a locale specifying a power operation on arbitrary
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   134
  types:
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   135
\<close>
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   136
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   137
locale %quote power =
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   138
  fixes power :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   139
  assumes power_commute: "power x \<circ> power y = power y \<circ> power x"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   140
begin
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   141
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   142
text \<open>
40752
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   143
  \noindent Inside that locale we can lift @{text power} to exponent
61779
9ace5e8310dc consolidated documentation
haftmann
parents: 61671
diff changeset
   144
  lists by means of a specification relative to that locale:
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   145
\<close>
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   146
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   147
primrec %quote powers :: "'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   148
  "powers [] = id"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   149
| "powers (x # xs) = power x \<circ> powers xs"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   150
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   151
lemma %quote powers_append:
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   152
  "powers (xs @ ys) = powers xs \<circ> powers ys"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   153
  by (induct xs) simp_all
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   154
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   155
lemma %quote powers_power:
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   156
  "powers xs \<circ> power x = power x \<circ> powers xs"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   157
  by (induct xs)
49739
13aa6d8268ec consolidated names of theorems on composition;
haftmann
parents: 48985
diff changeset
   158
    (simp_all del: o_apply id_apply add: comp_assoc,
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   159
      simp del: o_apply add: o_assoc power_commute)
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   160
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   161
lemma %quote powers_rev:
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   162
  "powers (rev xs) = powers xs"
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   163
    by (induct xs) (simp_all add: powers_append powers_power)
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   164
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   165
end %quote
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   166
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   167
text \<open>
38505
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38437
diff changeset
   168
  After an interpretation of this locale (say, @{command_def
61891
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61779
diff changeset
   169
  global_interpretation} @{text "fun_power:"} @{term [source] "power (\<lambda>n (f
61779
9ace5e8310dc consolidated documentation
haftmann
parents: 61671
diff changeset
   170
  :: 'a \<Rightarrow> 'a). f ^^ n)"}), one could naively expect to have a constant @{text
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   171
  "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
   172
  can be generated.  But this not the case: internally, the term
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   173
  @{text "fun_power.powers"} is an abbreviation for the foundational
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   174
  term @{term [source] "power.powers (\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)"}
58620
7435b6a3f72e more antiquotations;
wenzelm
parents: 57512
diff changeset
   175
  (see @{cite "isabelle-locale"} for the details behind).
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   176
61891
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61779
diff changeset
   177
  Fortunately, a succint solution is available: a dedicated
61779
9ace5e8310dc consolidated documentation
haftmann
parents: 61671
diff changeset
   178
  rewrite definition:
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   179
\<close>
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   180
61891
76189756ff65 documentation on last state of the art concerning interpretation
haftmann
parents: 61779
diff changeset
   181
global_interpretation %quote fun_power: power "(\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)"
61779
9ace5e8310dc consolidated documentation
haftmann
parents: 61671
diff changeset
   182
  defines funpows = fun_power.powers
9ace5e8310dc consolidated documentation
haftmann
parents: 61671
diff changeset
   183
  by unfold_locales
9ace5e8310dc consolidated documentation
haftmann
parents: 61671
diff changeset
   184
    (simp_all add: fun_eq_iff funpow_mult mult.commute)
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   185
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   186
text \<open>
61779
9ace5e8310dc consolidated documentation
haftmann
parents: 61671
diff changeset
   187
  \noindent This amends the interpretation morphisms such that
9ace5e8310dc consolidated documentation
haftmann
parents: 61671
diff changeset
   188
  occurrences of the foundational term @{term [source] "power.powers (\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)"}
9ace5e8310dc consolidated documentation
haftmann
parents: 61671
diff changeset
   189
  are folded to a newly defined constant @{const funpows}.
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   190
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   191
  After this setup procedure, code generation can continue as usual:
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   192
\<close>
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   193
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   194
text %quotetypewriter \<open>
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39664
diff changeset
   195
  @{code_stmts funpows (consts) Nat.funpow funpows (Haskell)}
61779
9ace5e8310dc consolidated documentation
haftmann
parents: 61671
diff changeset
   196
\<close>
37426
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   197
04d58897bf90 subsection on locale interpretation
haftmann
parents: 37210
diff changeset
   198
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   199
subsection \<open>Parallel computation\<close>
51172
16eb76ca1e4a note on parallel computation
haftmann
parents: 49739
diff changeset
   200
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   201
text \<open>
63680
6e1e8b5abbfa more symbols;
wenzelm
parents: 63669
diff changeset
   202
  Theory @{text Parallel} in \<^dir>\<open>~~/src/HOL/Library\<close> contains
51172
16eb76ca1e4a note on parallel computation
haftmann
parents: 49739
diff changeset
   203
  operations to exploit parallelism inside the Isabelle/ML
16eb76ca1e4a note on parallel computation
haftmann
parents: 49739
diff changeset
   204
  runtime engine.
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   205
\<close>
51172
16eb76ca1e4a note on parallel computation
haftmann
parents: 49739
diff changeset
   206
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   207
subsection \<open>Imperative data structures\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   208
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   209
text \<open>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   210
  If you consider imperative data structures as inevitable for a
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   211
  specific application, you should consider \emph{Imperative
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   212
  Functional Programming with Isabelle/HOL}
58620
7435b6a3f72e more antiquotations;
wenzelm
parents: 57512
diff changeset
   213
  @{cite "bulwahn-et-al:2008:imperative"}; the framework described there
40752
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   214
  is available in session @{text Imperative_HOL}, together with a
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   215
  short primer document.
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   216
\<close>
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   217
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   218
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   219
subsection \<open>ML system interfaces \label{sec:ml}\<close>
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   220
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   221
text \<open>
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   222
  Since the code generator framework not only aims to provide a nice
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   223
  Isar interface but also to form a base for code-generation-based
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   224
  applications, here a short description of the most fundamental ML
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   225
  interfaces.
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   226
\<close>
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   227
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   228
subsubsection \<open>Managing executable content\<close>
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   229
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   230
text %mlref \<open>
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   231
  \begin{mldecls}
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   232
  @{index_ML Code.read_const: "theory -> string -> string"} \\
63239
d562c9948dee explicit tagging of code equations de-baroquifies interface
haftmann
parents: 61891
diff changeset
   233
  @{index_ML Code.add_eqn: "Code.kind * bool -> thm -> theory -> theory"} \\
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   234
  @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51172
diff changeset
   235
  @{index_ML Code_Preproc.map_pre: "(Proof.context -> Proof.context) -> theory -> theory"} \\
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51172
diff changeset
   236
  @{index_ML Code_Preproc.map_post: "(Proof.context -> Proof.context) -> theory -> theory"} \\
38507
06728ef076a7 output whitespace tuning
haftmann
parents: 38505
diff changeset
   237
  @{index_ML Code_Preproc.add_functrans: "
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 55372
diff changeset
   238
    string * (Proof.context -> (thm * bool) list -> (thm * bool) list option)
38507
06728ef076a7 output whitespace tuning
haftmann
parents: 38505
diff changeset
   239
      -> theory -> theory"} \\
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   240
  @{index_ML Code_Preproc.del_functrans: "string -> theory -> theory"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   241
  @{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   242
  @{index_ML Code.get_type: "theory -> string
40752
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   243
    -> ((string * sort) list * (string * ((string * sort) list * typ list)) list) * bool"} \\
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   244
  @{index_ML Code.get_type_of_constr_or_abstr: "theory -> string -> (string * bool) option"}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   245
  \end{mldecls}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   246
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   247
  \begin{description}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   248
38510
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   249
  \item @{ML Code.read_const}~@{text thy}~@{text s}
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   250
     reads a constant as a concrete term expression @{text s}.
ec0408c7328b stub for evaluation chapter
haftmann
parents: 38507
diff changeset
   251
63239
d562c9948dee explicit tagging of code equations de-baroquifies interface
haftmann
parents: 61891
diff changeset
   252
  \item @{ML Code.add_eqn}~@{text "(kind, default)"}~@{text "thm"}~@{text "thy"} adds code equation
d562c9948dee explicit tagging of code equations de-baroquifies interface
haftmann
parents: 61891
diff changeset
   253
     @{text "thm"} to executable content.
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   254
63239
d562c9948dee explicit tagging of code equations de-baroquifies interface
haftmann
parents: 61891
diff changeset
   255
  \item @{ML Code.del_eqn}~@{text "thm"}~@{text "thy"} removes code equation
d562c9948dee explicit tagging of code equations de-baroquifies interface
haftmann
parents: 61891
diff changeset
   256
     @{text "thm"} from executable content, if present.
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   257
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   258
  \item @{ML Code_Preproc.map_pre}~@{text "f"}~@{text "thy"} changes
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   259
     the preprocessor simpset.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   260
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   261
  \item @{ML Code_Preproc.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   262
     function transformer @{text f} (named @{text name}) to executable content;
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   263
     @{text f} is a transformer of the code equations belonging
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   264
     to a certain function definition, depending on the
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   265
     current theory context.  Returning @{text NONE} indicates that no
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   266
     transformation took place;  otherwise, the whole process will be iterated
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   267
     with the new code equations.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   268
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   269
  \item @{ML Code_Preproc.del_functrans}~@{text "name"}~@{text "thy"} removes
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   270
     function transformer named @{text name} from executable content.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   271
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   272
  \item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   273
     a datatype to executable content, with generation
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   274
     set @{text cs}.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   275
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   276
  \item @{ML Code.get_type_of_constr_or_abstr}~@{text "thy"}~@{text "const"}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   277
     returns type constructor corresponding to
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   278
     constructor @{text const}; returns @{text NONE}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   279
     if @{text const} is no constructor.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   280
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   281
  \end{description}
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   282
\<close>
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   283
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   284
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   285
subsubsection \<open>Data depending on the theory's executable content\<close>
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   286
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   287
text \<open>
40752
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   288
  Implementing code generator applications on top of the framework set
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   289
  out so far usually not only involves using those primitive
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   290
  interfaces but also storing code-dependent data and various other
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   291
  things.
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   292
40752
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   293
  Due to incrementality of code generation, changes in the theory's
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   294
  executable content have to be propagated in a certain fashion.
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   295
  Additionally, such changes may occur not only during theory
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   296
  extension but also during theory merge, which is a little bit nasty
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   297
  from an implementation point of view.  The framework provides a
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   298
  solution to this technical challenge by providing a functorial data
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   299
  slot @{ML_functor Code_Data}; on instantiation of this functor, the
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   300
  following types and operations are required:
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   301
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   302
  \medskip
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   303
  \begin{tabular}{l}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   304
  @{text "type T"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   305
  @{text "val empty: T"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   306
  \end{tabular}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   307
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   308
  \begin{description}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   309
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   310
  \item @{text T} the type of data to store.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   311
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   312
  \item @{text empty} initial (empty) data.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   313
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   314
  \end{description}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   315
40752
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   316
  \noindent An instance of @{ML_functor Code_Data} provides the
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   317
  following interface:
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   318
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   319
  \medskip
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   320
  \begin{tabular}{l}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   321
  @{text "change: theory \<rightarrow> (T \<rightarrow> T) \<rightarrow> T"} \\
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   322
  @{text "change_yield: theory \<rightarrow> (T \<rightarrow> 'a * T) \<rightarrow> 'a * T"}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   323
  \end{tabular}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   324
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   325
  \begin{description}
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   326
40752
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   327
  \item @{text change} update of current data (cached!) by giving a
aae9a020fa77 tuned formatting;
haftmann
parents: 40352
diff changeset
   328
    continuation.
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   329
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   330
  \item @{text change_yield} update with side result.
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   331
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   332
  \end{description}
59377
056945909f60 modernized cartouches
haftmann
parents: 59003
diff changeset
   333
\<close>
38405
7935b334893e sketch of new outline
haftmann
parents: 37836
diff changeset
   334
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   335
end