src/Doc/Codegen/Adaptation.thy
author wenzelm
Mon, 15 Jul 2013 20:13:30 +0200
changeset 52665 5f817bad850a
parent 52378 08dbf9ff2140
child 52742 e7296939fec2
permissions -rw-r--r--
prefer @{file} references that are actually checked;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31050
555b56b66fcf adaptation replaces adaption
haftmann
parents: 30882
diff changeset
     1
theory Adaptation
28213
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
40351
090dac52cfd7 SMLdummy target
haftmann
parents: 39745
diff changeset
     5
setup %invisible {* Code_Target.extend_target ("\<SML>", ("SML", K I))
090dac52cfd7 SMLdummy target
haftmann
parents: 39745
diff changeset
     6
  #> Code_Target.extend_target ("\<SMLdummy>", ("Haskell", K I)) *}
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
     7
31050
555b56b66fcf adaptation replaces adaption
haftmann
parents: 30882
diff changeset
     8
section {* Adaptation to target languages \label{sec:adaptation} *}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
     9
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    10
subsection {* Adapting code generation *}
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    11
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    12
text {*
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    13
  The aspects of code generation introduced so far have two aspects
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    14
  in common:
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    15
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    16
  \begin{itemize}
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    17
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    18
    \item They act uniformly, without reference to a specific target
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    19
       language.
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    20
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    21
    \item They are \emph{safe} in the sense that as long as you trust
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    22
       the code generator meta theory and implementation, you cannot
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    23
       produce programs that yield results which are not derivable in
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    24
       the logic.
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    25
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    26
  \end{itemize}
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    27
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    28
  \noindent In this section we will introduce means to \emph{adapt}
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    29
  the serialiser to a specific target language, i.e.~to print program
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    30
  fragments in a way which accommodates \qt{already existing}
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    31
  ingredients of a target language environment, for three reasons:
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    32
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    33
  \begin{itemize}
28593
f087237af65d continued codegen tutorial
haftmann
parents: 28564
diff changeset
    34
    \item improving readability and aesthetics of generated code
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    35
    \item gaining efficiency
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    36
    \item interface with language parts which have no direct counterpart
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    37
      in @{text "HOL"} (say, imperative data structures)
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    38
  \end{itemize}
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    39
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    40
  \noindent Generally, you should avoid using those features yourself
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    41
  \emph{at any cost}:
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    42
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    43
  \begin{itemize}
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    44
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    45
    \item The safe configuration methods act uniformly on every target
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    46
      language, whereas for adaptation you have to treat each target
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    47
      language separately.
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    48
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    49
    \item Application is extremely tedious since there is no
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    50
      abstraction which would allow for a static check, making it easy
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    51
      to produce garbage.
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    52
34155
14aaccb399b3 Polishing up the English
paulson
parents: 31206
diff changeset
    53
    \item Subtle errors can be introduced unconsciously.
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    54
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    55
  \end{itemize}
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    56
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    57
  \noindent However, even if you ought refrain from setting up
51162
310b94ed1815 dropped now obsolete hint;
haftmann
parents: 51160
diff changeset
    58
  adaptation yourself, already @{text "HOL"} comes with some
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    59
  reasonable default adaptations (say, using target language list
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    60
  syntax).  There also some common adaptation cases which you can
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    61
  setup by importing particular library theories.  In order to
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    62
  understand these, we provide some clues here; these however are not
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    63
  supposed to replace a careful study of the sources.
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    64
*}
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    65
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    66
31050
555b56b66fcf adaptation replaces adaption
haftmann
parents: 30882
diff changeset
    67
subsection {* The adaptation principle *}
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    68
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
    69
text {*
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    70
  Figure \ref{fig:adaptation} illustrates what \qt{adaptation} is
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    71
  conceptually supposed to be:
28601
b72589374396 figure for adaption
haftmann
parents: 28593
diff changeset
    72
b72589374396 figure for adaption
haftmann
parents: 28593
diff changeset
    73
  \begin{figure}[here]
48954
c548d26daa8c avoid clash with generated Adaptation.tex on case-insensible file-systems;
wenzelm
parents: 48951
diff changeset
    74
    \includegraphics{adapt}
31050
555b56b66fcf adaptation replaces adaption
haftmann
parents: 30882
diff changeset
    75
    \caption{The adaptation principle}
555b56b66fcf adaptation replaces adaption
haftmann
parents: 30882
diff changeset
    76
    \label{fig:adaptation}
28601
b72589374396 figure for adaption
haftmann
parents: 28593
diff changeset
    77
  \end{figure}
b72589374396 figure for adaption
haftmann
parents: 28593
diff changeset
    78
b72589374396 figure for adaption
haftmann
parents: 28593
diff changeset
    79
  \noindent In the tame view, code generation acts as broker between
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    80
  @{text logic}, @{text "intermediate language"} and @{text "target
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    81
  language"} by means of @{text translation} and @{text
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    82
  serialisation}; for the latter, the serialiser has to observe the
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    83
  structure of the @{text language} itself plus some @{text reserved}
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    84
  keywords which have to be avoided for generated code.  However, if
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    85
  you consider @{text adaptation} mechanisms, the code generated by
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    86
  the serializer is just the tip of the iceberg:
28601
b72589374396 figure for adaption
haftmann
parents: 28593
diff changeset
    87
b72589374396 figure for adaption
haftmann
parents: 28593
diff changeset
    88
  \begin{itemize}
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    89
28635
cc53d2ab0170 filled remaining gaps
haftmann
parents: 28609
diff changeset
    90
    \item @{text serialisation} can be \emph{parametrised} such that
cc53d2ab0170 filled remaining gaps
haftmann
parents: 28609
diff changeset
    91
      logical entities are mapped to target-specific ones
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    92
      (e.g. target-specific list syntax, see also
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    93
      \secref{sec:adaptation_mechanisms})
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    94
28635
cc53d2ab0170 filled remaining gaps
haftmann
parents: 28609
diff changeset
    95
    \item Such parametrisations can involve references to a
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    96
      target-specific standard @{text library} (e.g. using the @{text
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    97
      Haskell} @{verbatim Maybe} type instead of the @{text HOL}
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    98
      @{type "option"} type); if such are used, the corresponding
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
    99
      identifiers (in our example, @{verbatim Maybe}, @{verbatim
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   100
      Nothing} and @{verbatim Just}) also have to be considered @{text
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   101
      reserved}.
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   102
28635
cc53d2ab0170 filled remaining gaps
haftmann
parents: 28609
diff changeset
   103
    \item Even more, the user can enrich the library of the
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   104
      target-language by providing code snippets (\qt{@{text
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   105
      "includes"}}) which are prepended to any generated code (see
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   106
      \secref{sec:include}); this typically also involves further
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   107
      @{text reserved} identifiers.
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   108
28601
b72589374396 figure for adaption
haftmann
parents: 28593
diff changeset
   109
  \end{itemize}
28635
cc53d2ab0170 filled remaining gaps
haftmann
parents: 28609
diff changeset
   110
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   111
  \noindent As figure \ref{fig:adaptation} illustrates, all these
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   112
  adaptation mechanisms have to act consistently; it is at the
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   113
  discretion of the user to take care for this.
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   114
*}
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   115
31050
555b56b66fcf adaptation replaces adaption
haftmann
parents: 30882
diff changeset
   116
subsection {* Common adaptation patterns *}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   117
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   118
text {*
28428
haftmann
parents: 28420
diff changeset
   119
  The @{theory HOL} @{theory Main} theory already provides a code
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   120
  generator setup which should be suitable for most applications.
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   121
  Common extensions and modifications are available by certain
52665
5f817bad850a prefer @{file} references that are actually checked;
wenzelm
parents: 52378
diff changeset
   122
  theories in @{file "~~/src/HOL/Library"}; beside being useful in
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   123
  applications, they may serve as a tutorial for customising the code
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   124
  generator setup (see below \secref{sec:adaptation_mechanisms}).
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   125
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   126
  \begin{description}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   127
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   128
    \item[@{theory "Code_Numeral"}] provides additional numeric
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   129
       types @{typ integer} and @{typ natural} isomorphic to types
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   130
       @{typ int} and @{typ nat} respectively.  Type @{typ integer}
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   131
       is mapped to target-language built-in integers; @{typ natural}
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   132
       is implemented as abstract type over @{typ integer}.
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   133
       Useful for code setups which involve e.g.~indexing
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   134
       of target-language arrays.  Part of @{text "HOL-Main"}.
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   135
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   136
    \item[@{text "Code_Target_Int"}] implements type @{typ int}
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   137
       by @{typ integer} and thus by target-language built-in integers.
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   138
51171
e8b2d90da499 corrected and clarified Code_Binary_Nat vs. Code_Target_Nat
haftmann
parents: 51162
diff changeset
   139
    \item[@{text "Code_Binary_Nat"}] implements type
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   140
       @{typ nat} using a binary rather than a linear representation,
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   141
       which yields a considerable speedup for computations.
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   142
       Pattern matching with @{term "0\<Colon>nat"} / @{const "Suc"} is eliminated
51171
e8b2d90da499 corrected and clarified Code_Binary_Nat vs. Code_Target_Nat
haftmann
parents: 51162
diff changeset
   143
       by a preprocessor.\label{abstract_nat}
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   144
51171
e8b2d90da499 corrected and clarified Code_Binary_Nat vs. Code_Target_Nat
haftmann
parents: 51162
diff changeset
   145
    \item[@{text "Code_Target_Nat"}] implements type @{typ nat}
e8b2d90da499 corrected and clarified Code_Binary_Nat vs. Code_Target_Nat
haftmann
parents: 51162
diff changeset
   146
       by @{typ integer} and thus by target-language built-in integers.
e8b2d90da499 corrected and clarified Code_Binary_Nat vs. Code_Target_Nat
haftmann
parents: 51162
diff changeset
   147
       Pattern matching with @{term "0\<Colon>nat"} / @{const "Suc"} is eliminated
e8b2d90da499 corrected and clarified Code_Binary_Nat vs. Code_Target_Nat
haftmann
parents: 51162
diff changeset
   148
       by a preprocessor.
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   149
51162
310b94ed1815 dropped now obsolete hint;
haftmann
parents: 51160
diff changeset
   150
    \item[@{text "Code_Target_Numeral"}] is a convenience theory
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   151
       containing both @{text "Code_Target_Nat"} and
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   152
       @{text "Code_Target_Int"}.
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   153
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   154
    \item[@{text "Code_Char"}] represents @{text HOL} characters by
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
   155
       character literals in target languages.
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   156
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   157
    \item[@{theory "String"}] provides an additional datatype @{typ
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   158
       String.literal} which is isomorphic to strings; @{typ
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   159
       String.literal}s are mapped to target-language strings.  Useful
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   160
       for code setups which involve e.g.~printing (error) messages.
46519
17dde5feea4b clarified
haftmann
parents: 40351
diff changeset
   161
       Part of @{text "HOL-Main"}.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   162
51162
310b94ed1815 dropped now obsolete hint;
haftmann
parents: 51160
diff changeset
   163
    \item[@{theory "IArray"}] provides a type @{typ "'a iarray"}
310b94ed1815 dropped now obsolete hint;
haftmann
parents: 51160
diff changeset
   164
       isomorphic to lists but implemented by (effectively immutable)
310b94ed1815 dropped now obsolete hint;
haftmann
parents: 51160
diff changeset
   165
       arrays \emph{in SML only}.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   166
51162
310b94ed1815 dropped now obsolete hint;
haftmann
parents: 51160
diff changeset
   167
  \end{description}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   168
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   169
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   170
31050
555b56b66fcf adaptation replaces adaption
haftmann
parents: 30882
diff changeset
   171
subsection {* Parametrising serialisation \label{sec:adaptation_mechanisms} *}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   172
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   173
text {*
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   174
  Consider the following function and its corresponding SML code:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   175
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   176
28564
haftmann
parents: 28561
diff changeset
   177
primrec %quote in_interval :: "nat \<times> nat \<Rightarrow> nat \<Rightarrow> bool" where
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   178
  "in_interval (k, l) n \<longleftrightarrow> k \<le> n \<and> n \<le> l"
28447
haftmann
parents: 28428
diff changeset
   179
(*<*)
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   180
code_printing %invisible
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   181
  type_constructor bool \<rightharpoonup> (SML)
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   182
| constant True \<rightharpoonup> (SML)
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   183
| constant False \<rightharpoonup> (SML)
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   184
| constant HOL.conj \<rightharpoonup> (SML)
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   185
| constant Not \<rightharpoonup> (SML)
28447
haftmann
parents: 28428
diff changeset
   186
(*>*)
39745
3aa2bc9c5478 combine quote and typewriter/tt tag
haftmann
parents: 39711
diff changeset
   187
text %quotetypewriter {*
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39664
diff changeset
   188
  @{code_stmts in_interval (SML)}
39664
0afaf89ab591 more canonical type setting of type writer code examples
haftmann
parents: 39643
diff changeset
   189
*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   190
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   191
text {*
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   192
  \noindent Though this is correct code, it is a little bit
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   193
  unsatisfactory: boolean values and operators are materialised as
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   194
  distinguished entities with have nothing to do with the SML-built-in
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   195
  notion of \qt{bool}.  This results in less readable code;
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   196
  additionally, eager evaluation may cause programs to loop or break
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   197
  which would perfectly terminate when the existing SML @{verbatim
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   198
  "bool"} would be used.  To map the HOL @{typ bool} on SML @{verbatim
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   199
  "bool"}, we may use \qn{custom serialisations}:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   200
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   201
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   202
code_printing %quotett
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   203
  type_constructor bool \<rightharpoonup> (SML) "bool"
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   204
| constant True \<rightharpoonup> (SML) "true"
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   205
| constant False \<rightharpoonup> (SML) "false"
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   206
| constant HOL.conj \<rightharpoonup> (SML) "_ andalso _"
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   207
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   208
text {*
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   209
  \noindent The @{command_def code_printing} command takes a series
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   210
  of symbols (contants, type constructor, \ldots)
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   211
  together with target-specific custom serialisations.  Each
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   212
  custom serialisation starts with a target language identifier
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   213
  followed by an expression, which during code serialisation is
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   214
  inserted whenever the type constructor would occur.  Each
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   215
  ``@{verbatim "_"}'' in a serialisation expression is treated as a
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   216
  placeholder for the constant's or the type constructor's arguments.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   217
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   218
39745
3aa2bc9c5478 combine quote and typewriter/tt tag
haftmann
parents: 39711
diff changeset
   219
text %quotetypewriter {*
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39664
diff changeset
   220
  @{code_stmts in_interval (SML)}
39664
0afaf89ab591 more canonical type setting of type writer code examples
haftmann
parents: 39643
diff changeset
   221
*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   222
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   223
text {*
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   224
  \noindent This still is not perfect: the parentheses around the
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   225
  \qt{andalso} expression are superfluous.  Though the serialiser by
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   226
  no means attempts to imitate the rich Isabelle syntax framework, it
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   227
  provides some common idioms, notably associative infixes with
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   228
  precedences which may be used here:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   229
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   230
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   231
code_printing %quotett
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   232
  constant HOL.conj \<rightharpoonup> (SML) infixl 1 "andalso"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   233
39745
3aa2bc9c5478 combine quote and typewriter/tt tag
haftmann
parents: 39711
diff changeset
   234
text %quotetypewriter {*
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39664
diff changeset
   235
  @{code_stmts in_interval (SML)}
39664
0afaf89ab591 more canonical type setting of type writer code examples
haftmann
parents: 39643
diff changeset
   236
*}
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   237
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   238
text {*
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   239
  \noindent The attentive reader may ask how we assert that no
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   240
  generated code will accidentally overwrite.  For this reason the
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   241
  serialiser has an internal table of identifiers which have to be
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   242
  avoided to be used for new declarations.  Initially, this table
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   243
  typically contains the keywords of the target language.  It can be
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   244
  extended manually, thus avoiding accidental overwrites, using the
38505
2f8699695cf6 use command_def vs. command more consciously
haftmann
parents: 38450
diff changeset
   245
  @{command_def "code_reserved"} command:
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   246
*}
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   247
40351
090dac52cfd7 SMLdummy target
haftmann
parents: 39745
diff changeset
   248
code_reserved %quote "\<SMLdummy>" bool true false andalso
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   249
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   250
text {*
28447
haftmann
parents: 28428
diff changeset
   251
  \noindent Next, we try to map HOL pairs to SML pairs, using the
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   252
  infix ``@{verbatim "*"}'' type constructor and parentheses:
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   253
*}
28447
haftmann
parents: 28428
diff changeset
   254
(*<*)
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   255
code_printing %invisible
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   256
  type_constructor prod \<rightharpoonup> (SML)
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   257
| constant Pair \<rightharpoonup> (SML)
28447
haftmann
parents: 28428
diff changeset
   258
(*>*)
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   259
code_printing %quotett
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   260
  type_constructor prod \<rightharpoonup> (SML) infix 2 "*"
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   261
| constant Pair \<rightharpoonup> (SML) "!((_),/ (_))"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   262
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   263
text {*
28593
f087237af65d continued codegen tutorial
haftmann
parents: 28564
diff changeset
   264
  \noindent The initial bang ``@{verbatim "!"}'' tells the serialiser
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   265
  never to put parentheses around the whole expression (they are
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   266
  already present), while the parentheses around argument place
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   267
  holders tell not to put parentheses around the arguments.  The slash
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   268
  ``@{verbatim "/"}'' (followed by arbitrary white space) inserts a
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   269
  space which may be used as a break if necessary during pretty
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   270
  printing.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   271
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   272
  These examples give a glimpse what mechanisms custom serialisations
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   273
  provide; however their usage requires careful thinking in order not
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   274
  to introduce inconsistencies -- or, in other words: custom
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   275
  serialisations are completely axiomatic.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   276
39643
29cc021398fc corrections and tuning
haftmann
parents: 39063
diff changeset
   277
  A further noteworthy detail is that any special character in a
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   278
  custom serialisation may be quoted using ``@{verbatim "'"}''; thus,
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   279
  in ``@{verbatim "fn '_ => _"}'' the first ``@{verbatim "_"}'' is a
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   280
  proper underscore while the second ``@{verbatim "_"}'' is a
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   281
  placeholder.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   282
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   283
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   284
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   285
subsection {* @{text Haskell} serialisation *}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   286
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   287
text {*
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   288
  For convenience, the default @{text HOL} setup for @{text Haskell}
39063
5f9692dd621f adapted to change eq -> equal
haftmann
parents: 38506
diff changeset
   289
  maps the @{class equal} class to its counterpart in @{text Haskell},
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   290
  giving custom serialisations for the class @{class equal}
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   291
  and its operation @{const [source] HOL.equal}.
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   292
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   293
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   294
code_printing %quotett
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   295
  type_class equal \<rightharpoonup> (Haskell) "Eq"
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   296
| constant HOL.equal \<rightharpoonup> (Haskell) infixl 4 "=="
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   297
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   298
text {*
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   299
  \noindent A problem now occurs whenever a type which is an instance
39063
5f9692dd621f adapted to change eq -> equal
haftmann
parents: 38506
diff changeset
   300
  of @{class equal} in @{text HOL} is mapped on a @{text
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   301
  Haskell}-built-in type which is also an instance of @{text Haskell}
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   302
  @{text Eq}:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   303
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   304
28564
haftmann
parents: 28561
diff changeset
   305
typedecl %quote bar
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   306
39063
5f9692dd621f adapted to change eq -> equal
haftmann
parents: 38506
diff changeset
   307
instantiation %quote bar :: equal
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   308
begin
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   309
39063
5f9692dd621f adapted to change eq -> equal
haftmann
parents: 38506
diff changeset
   310
definition %quote "HOL.equal (x\<Colon>bar) y \<longleftrightarrow> x = y"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   311
39063
5f9692dd621f adapted to change eq -> equal
haftmann
parents: 38506
diff changeset
   312
instance %quote by default (simp add: equal_bar_def)
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
   313
30880
257cbe43faa8 tuned manual
haftmann
parents: 30836
diff changeset
   314
end %quote (*<*)
257cbe43faa8 tuned manual
haftmann
parents: 30836
diff changeset
   315
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   316
(*>*) code_printing %quotett
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   317
  type_constructor bar \<rightharpoonup> (Haskell) "Integer"
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   318
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   319
text {*
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   320
  \noindent The code generator would produce an additional instance,
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   321
  which of course is rejected by the @{text Haskell} compiler.  To
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   322
  suppress this additional instance:
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   323
*}
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   324
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   325
code_printing %quotett
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   326
  class_instance bar :: "HOL.equal" \<rightharpoonup> (Haskell) -
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   327
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   328
28635
cc53d2ab0170 filled remaining gaps
haftmann
parents: 28609
diff changeset
   329
subsection {* Enhancing the target language context \label{sec:include} *}
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   330
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   331
text {*
28593
f087237af65d continued codegen tutorial
haftmann
parents: 28564
diff changeset
   332
  In rare cases it is necessary to \emph{enrich} the context of a
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   333
  target language; this can also be accomplished using the @{command
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   334
  "code_printing"} command:
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   335
*}
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   336
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   337
code_printing %quotett
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   338
  code_module "Errno" \<rightharpoonup> (Haskell) {*errno i = error ("Error number: " ++ show i)*}
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   339
39745
3aa2bc9c5478 combine quote and typewriter/tt tag
haftmann
parents: 39711
diff changeset
   340
code_reserved %quotett Haskell Errno
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   341
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   342
text {*
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   343
  \noindent Such named modules are then prepended to every
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   344
  generated code.  Inspect such code in order to find out how
52378
08dbf9ff2140 documentation on code_printing and code_identifier
haftmann
parents: 51172
diff changeset
   345
  this behaves with respect to a particular
38450
ada5814c9d87 tuned whitespace
haftmann
parents: 37836
diff changeset
   346
  target language.
28561
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   347
*}
056255ade52a some adaption
haftmann
parents: 28456
diff changeset
   348
28419
f65e8b318581 re-canibalised manual
haftmann
parents: 28213
diff changeset
   349
end
46519
17dde5feea4b clarified
haftmann
parents: 40351
diff changeset
   350