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