| author | wenzelm | 
| Tue, 07 Oct 2014 22:35:11 +0200 | |
| changeset 58620 | 7435b6a3f72e | 
| parent 55147 | bce3dbc11f95 | 
| child 59104 | a14475f044b2 | 
| permissions | -rw-r--r-- | 
| 31050 | 1  | 
theory Adaptation  | 
| 28213 | 2  | 
imports Setup  | 
3  | 
begin  | 
|
4  | 
||
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
52742 
diff
changeset
 | 
5  | 
setup %invisible {* Code_Target.extend_target ("\<SML>", ("SML", I))
 | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
52742 
diff
changeset
 | 
6  | 
  #> Code_Target.extend_target ("\<SMLdummy>", ("Haskell", I)) *}
 | 
| 28561 | 7  | 
|
| 31050 | 8  | 
section {* Adaptation to target languages \label{sec:adaptation} *}
 | 
| 28419 | 9  | 
|
| 28561 | 10  | 
subsection {* Adapting code generation *}
 | 
11  | 
||
12  | 
text {*
 | 
|
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.  | 
|
| 28561 | 64  | 
*}  | 
65  | 
||
| 38450 | 66  | 
|
| 31050 | 67  | 
subsection {* The adaptation principle *}
 | 
| 28561 | 68  | 
|
69  | 
text {*
 | 
|
| 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.  | 
|
| 28561 | 149  | 
*}  | 
150  | 
||
| 31050 | 151  | 
subsection {* Common adaptation patterns *}
 | 
| 28419 | 152  | 
|
153  | 
text {*
 | 
|
| 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: 
52378 
diff
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: 
48985 
diff
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: 
48985 
diff
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: 
48985 
diff
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: 
48985 
diff
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: 
48985 
diff
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: 
48985 
diff
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: 
48985 
diff
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: 
48985 
diff
changeset
 | 
170  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48985 
diff
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: 
48985 
diff
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: 
51162 
diff
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: 
48985 
diff
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: 
48985 
diff
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: 
48985 
diff
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: 
51162 
diff
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: 
48985 
diff
changeset
 | 
179  | 
|
| 
51171
 
e8b2d90da499
corrected and clarified Code_Binary_Nat vs. Code_Target_Nat
 
haftmann 
parents: 
51162 
diff
changeset
 | 
180  | 
    \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
 | 
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: 
51162 
diff
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: 
51162 
diff
changeset
 | 
183  | 
by a preprocessor.  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48985 
diff
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: 
48985 
diff
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: 
48985 
diff
changeset
 | 
187  | 
       @{text "Code_Target_Int"}.
 | 
| 38450 | 188  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48985 
diff
changeset
 | 
189  | 
    \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
 | 
190  | 
character literals in target languages.  | 
| 38450 | 191  | 
|
192  | 
    \item[@{theory "String"}] provides an additional datatype @{typ
 | 
|
193  | 
       String.literal} which is isomorphic to strings; @{typ
 | 
|
194  | 
String.literal}s are mapped to target-language strings. Useful  | 
|
195  | 
for code setups which involve e.g.~printing (error) messages.  | 
|
| 46519 | 196  | 
       Part of @{text "HOL-Main"}.
 | 
| 28419 | 197  | 
|
| 51162 | 198  | 
    \item[@{theory "IArray"}] provides a type @{typ "'a iarray"}
 | 
199  | 
isomorphic to lists but implemented by (effectively immutable)  | 
|
200  | 
       arrays \emph{in SML only}.
 | 
|
| 28419 | 201  | 
|
| 51162 | 202  | 
  \end{description}
 | 
| 28419 | 203  | 
*}  | 
204  | 
||
205  | 
||
| 31050 | 206  | 
subsection {* Parametrising serialisation \label{sec:adaptation_mechanisms} *}
 | 
| 28419 | 207  | 
|
208  | 
text {*
 | 
|
| 38450 | 209  | 
Consider the following function and its corresponding SML code:  | 
| 28419 | 210  | 
*}  | 
211  | 
||
| 28564 | 212  | 
primrec %quote in_interval :: "nat \<times> nat \<Rightarrow> nat \<Rightarrow> bool" where  | 
| 28419 | 213  | 
"in_interval (k, l) n \<longleftrightarrow> k \<le> n \<and> n \<le> l"  | 
| 28447 | 214  | 
(*<*)  | 
| 
52378
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
215  | 
code_printing %invisible  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
216  | 
type_constructor bool \<rightharpoonup> (SML)  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
217  | 
| constant True \<rightharpoonup> (SML)  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
218  | 
| constant False \<rightharpoonup> (SML)  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
219  | 
| constant HOL.conj \<rightharpoonup> (SML)  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
220  | 
| constant Not \<rightharpoonup> (SML)  | 
| 28447 | 221  | 
(*>*)  | 
| 39745 | 222  | 
text %quotetypewriter {*
 | 
| 39683 | 223  | 
  @{code_stmts in_interval (SML)}
 | 
| 
39664
 
0afaf89ab591
more canonical type setting of type writer code examples
 
haftmann 
parents: 
39643 
diff
changeset
 | 
224  | 
*}  | 
| 28419 | 225  | 
|
226  | 
text {*
 | 
|
| 38450 | 227  | 
\noindent Though this is correct code, it is a little bit  | 
228  | 
unsatisfactory: boolean values and operators are materialised as  | 
|
229  | 
distinguished entities with have nothing to do with the SML-built-in  | 
|
230  | 
  notion of \qt{bool}.  This results in less readable code;
 | 
|
231  | 
additionally, eager evaluation may cause programs to loop or break  | 
|
232  | 
  which would perfectly terminate when the existing SML @{verbatim
 | 
|
233  | 
  "bool"} would be used.  To map the HOL @{typ bool} on SML @{verbatim
 | 
|
234  | 
  "bool"}, we may use \qn{custom serialisations}:
 | 
|
| 28419 | 235  | 
*}  | 
236  | 
||
| 
52378
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
237  | 
code_printing %quotett  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
238  | 
type_constructor bool \<rightharpoonup> (SML) "bool"  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
239  | 
| constant True \<rightharpoonup> (SML) "true"  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
240  | 
| constant False \<rightharpoonup> (SML) "false"  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
241  | 
| constant HOL.conj \<rightharpoonup> (SML) "_ andalso _"  | 
| 28213 | 242  | 
|
| 28419 | 243  | 
text {*
 | 
| 
52378
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
244  | 
  \noindent The @{command_def code_printing} command takes a series
 | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
245  | 
of symbols (contants, type constructor, \ldots)  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
246  | 
together with target-specific custom serialisations. Each  | 
| 38450 | 247  | 
custom serialisation starts with a target language identifier  | 
248  | 
followed by an expression, which during code serialisation is  | 
|
| 
52378
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
249  | 
inserted whenever the type constructor would occur. Each  | 
| 38450 | 250  | 
  ``@{verbatim "_"}'' in a serialisation expression is treated as a
 | 
| 
52378
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
251  | 
placeholder for the constant's or the type constructor's arguments.  | 
| 28419 | 252  | 
*}  | 
253  | 
||
| 39745 | 254  | 
text %quotetypewriter {*
 | 
| 39683 | 255  | 
  @{code_stmts in_interval (SML)}
 | 
| 
39664
 
0afaf89ab591
more canonical type setting of type writer code examples
 
haftmann 
parents: 
39643 
diff
changeset
 | 
256  | 
*}  | 
| 28419 | 257  | 
|
258  | 
text {*
 | 
|
| 38450 | 259  | 
\noindent This still is not perfect: the parentheses around the  | 
260  | 
  \qt{andalso} expression are superfluous.  Though the serialiser by
 | 
|
261  | 
no means attempts to imitate the rich Isabelle syntax framework, it  | 
|
262  | 
provides some common idioms, notably associative infixes with  | 
|
263  | 
precedences which may be used here:  | 
|
| 28419 | 264  | 
*}  | 
265  | 
||
| 
52378
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
266  | 
code_printing %quotett  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
267  | 
constant HOL.conj \<rightharpoonup> (SML) infixl 1 "andalso"  | 
| 28419 | 268  | 
|
| 39745 | 269  | 
text %quotetypewriter {*
 | 
| 39683 | 270  | 
  @{code_stmts in_interval (SML)}
 | 
| 
39664
 
0afaf89ab591
more canonical type setting of type writer code examples
 
haftmann 
parents: 
39643 
diff
changeset
 | 
271  | 
*}  | 
| 28419 | 272  | 
|
273  | 
text {*
 | 
|
| 38450 | 274  | 
\noindent The attentive reader may ask how we assert that no  | 
275  | 
generated code will accidentally overwrite. For this reason the  | 
|
276  | 
serialiser has an internal table of identifiers which have to be  | 
|
277  | 
avoided to be used for new declarations. Initially, this table  | 
|
278  | 
typically contains the keywords of the target language. It can be  | 
|
279  | 
extended manually, thus avoiding accidental overwrites, using the  | 
|
| 38505 | 280  | 
  @{command_def "code_reserved"} command:
 | 
| 28561 | 281  | 
*}  | 
282  | 
||
| 40351 | 283  | 
code_reserved %quote "\<SMLdummy>" bool true false andalso  | 
| 28561 | 284  | 
|
285  | 
text {*
 | 
|
| 28447 | 286  | 
\noindent Next, we try to map HOL pairs to SML pairs, using the  | 
| 28419 | 287  | 
  infix ``@{verbatim "*"}'' type constructor and parentheses:
 | 
288  | 
*}  | 
|
| 28447 | 289  | 
(*<*)  | 
| 
52378
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
290  | 
code_printing %invisible  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
291  | 
type_constructor prod \<rightharpoonup> (SML)  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
292  | 
| constant Pair \<rightharpoonup> (SML)  | 
| 28447 | 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_constructor prod \<rightharpoonup> (SML) infix 2 "*"  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
296  | 
| constant Pair \<rightharpoonup> (SML) "!((_),/ (_))"  | 
| 28419 | 297  | 
|
298  | 
text {*
 | 
|
| 28593 | 299  | 
  \noindent The initial bang ``@{verbatim "!"}'' tells the serialiser
 | 
| 38450 | 300  | 
never to put parentheses around the whole expression (they are  | 
301  | 
already present), while the parentheses around argument place  | 
|
302  | 
holders tell not to put parentheses around the arguments. The slash  | 
|
303  | 
  ``@{verbatim "/"}'' (followed by arbitrary white space) inserts a
 | 
|
304  | 
space which may be used as a break if necessary during pretty  | 
|
305  | 
printing.  | 
|
| 28419 | 306  | 
|
| 38450 | 307  | 
These examples give a glimpse what mechanisms custom serialisations  | 
308  | 
provide; however their usage requires careful thinking in order not  | 
|
309  | 
to introduce inconsistencies -- or, in other words: custom  | 
|
310  | 
serialisations are completely axiomatic.  | 
|
| 28419 | 311  | 
|
| 39643 | 312  | 
A further noteworthy detail is that any special character in a  | 
| 38450 | 313  | 
  custom serialisation may be quoted using ``@{verbatim "'"}''; thus,
 | 
314  | 
  in ``@{verbatim "fn '_ => _"}'' the first ``@{verbatim "_"}'' is a
 | 
|
315  | 
  proper underscore while the second ``@{verbatim "_"}'' is a
 | 
|
316  | 
placeholder.  | 
|
| 28419 | 317  | 
*}  | 
318  | 
||
319  | 
||
320  | 
subsection {* @{text Haskell} serialisation *}
 | 
|
321  | 
||
322  | 
text {*
 | 
|
| 38450 | 323  | 
  For convenience, the default @{text HOL} setup for @{text Haskell}
 | 
| 39063 | 324  | 
  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
 | 
325  | 
  giving custom serialisations for the class @{class equal}
 | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
326  | 
  and its operation @{const [source] HOL.equal}.
 | 
| 28419 | 327  | 
*}  | 
328  | 
||
| 
52378
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
329  | 
code_printing %quotett  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
330  | 
type_class equal \<rightharpoonup> (Haskell) "Eq"  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
331  | 
| constant HOL.equal \<rightharpoonup> (Haskell) infixl 4 "=="  | 
| 28419 | 332  | 
|
333  | 
text {*
 | 
|
| 38450 | 334  | 
\noindent A problem now occurs whenever a type which is an instance  | 
| 39063 | 335  | 
  of @{class equal} in @{text HOL} is mapped on a @{text
 | 
| 38450 | 336  | 
  Haskell}-built-in type which is also an instance of @{text Haskell}
 | 
337  | 
  @{text Eq}:
 | 
|
| 28419 | 338  | 
*}  | 
339  | 
||
| 28564 | 340  | 
typedecl %quote bar  | 
| 28419 | 341  | 
|
| 39063 | 342  | 
instantiation %quote bar :: equal  | 
| 28419 | 343  | 
begin  | 
344  | 
||
| 39063 | 345  | 
definition %quote "HOL.equal (x\<Colon>bar) y \<longleftrightarrow> x = y"  | 
| 28419 | 346  | 
|
| 39063 | 347  | 
instance %quote by default (simp add: equal_bar_def)  | 
| 28213 | 348  | 
|
| 30880 | 349  | 
end %quote (*<*)  | 
350  | 
||
| 
52378
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
351  | 
(*>*) code_printing %quotett  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
352  | 
type_constructor bar \<rightharpoonup> (Haskell) "Integer"  | 
| 28419 | 353  | 
|
354  | 
text {*
 | 
|
| 38450 | 355  | 
\noindent The code generator would produce an additional instance,  | 
356  | 
  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
 | 
357  | 
suppress this additional instance:  | 
| 28419 | 358  | 
*}  | 
359  | 
||
| 
52378
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
360  | 
code_printing %quotett  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
361  | 
class_instance bar :: "HOL.equal" \<rightharpoonup> (Haskell) -  | 
| 28419 | 362  | 
|
| 28561 | 363  | 
|
| 28635 | 364  | 
subsection {* Enhancing the target language context \label{sec:include} *}
 | 
| 28561 | 365  | 
|
366  | 
text {*
 | 
|
| 28593 | 367  | 
  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
 | 
368  | 
  target language; this can also be accomplished using the @{command
 | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
369  | 
"code_printing"} command:  | 
| 28561 | 370  | 
*}  | 
371  | 
||
| 
52378
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
372  | 
code_printing %quotett  | 
| 
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
373  | 
  code_module "Errno" \<rightharpoonup> (Haskell) {*errno i = error ("Error number: " ++ show i)*}
 | 
| 28561 | 374  | 
|
| 39745 | 375  | 
code_reserved %quotett Haskell Errno  | 
| 28561 | 376  | 
|
377  | 
text {*
 | 
|
| 
52378
 
08dbf9ff2140
documentation on code_printing and code_identifier
 
haftmann 
parents: 
51172 
diff
changeset
 | 
378  | 
\noindent Such named modules are then prepended to every  | 
| 38450 | 379  | 
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
 | 
380  | 
this behaves with respect to a particular  | 
| 38450 | 381  | 
target language.  | 
| 28561 | 382  | 
*}  | 
383  | 
||
| 28419 | 384  | 
end  | 
| 46519 | 385  |