5 section {* Introduction and Overview *} |
5 section {* Introduction and Overview *} |
6 |
6 |
7 text {* |
7 text {* |
8 This tutorial introduces a generic code generator for the |
8 This tutorial introduces a generic code generator for the |
9 @{text Isabelle} system. |
9 @{text Isabelle} system. |
10 Generic in the sense that the |
10 The |
11 \qn{target language} for which code shall ultimately be |
11 \qn{target language} for which code is |
12 generated is not fixed but may be an arbitrary state-of-the-art |
12 generated is not fixed, but may be one of several |
13 functional programming language (currently, the implementation |
13 functional programming languages (currently, the implementation |
14 supports @{text SML} \cite{SML}, @{text OCaml} \cite{OCaml} and @{text Haskell} |
14 supports @{text SML} \cite{SML}, @{text OCaml} \cite{OCaml} and @{text Haskell} |
15 \cite{haskell-revised-report}). |
15 \cite{haskell-revised-report}). |
16 |
16 |
17 Conceptually the code generator framework is part |
17 Conceptually the code generator framework is part |
18 of Isabelle's @{theory Pure} meta logic framework; the logic |
18 of Isabelle's @{theory Pure} meta logic framework; the logic |
19 @{theory HOL} which is an extension of @{theory Pure} |
19 @{theory HOL} \cite{isa-tutorial}, which is an extension of @{theory Pure}, |
20 already comes with a reasonable framework setup and thus provides |
20 already comes with a reasonable framework setup and thus provides |
21 a good working horse for raising code-generation-driven |
21 a good basis for creating code-generation-driven |
22 applications. So, we assume some familiarity and experience |
22 applications. So, we assume some familiarity and experience |
23 with the ingredients of the @{theory HOL} distribution theories. |
23 with the ingredients of the @{theory HOL} distribution theories. |
24 (see also \cite{isa-tutorial}). |
|
25 |
24 |
26 The code generator aims to be usable with no further ado |
25 The code generator aims to be usable with no further ado |
27 in most cases while allowing for detailed customisation. |
26 in most cases, while allowing for detailed customisation. |
28 This manifests in the structure of this tutorial: after a short |
27 This can be seen in the structure of this tutorial: after a short |
29 conceptual introduction with an example (\secref{sec:intro}), |
28 conceptual introduction with an example (\secref{sec:intro}), |
30 we discuss the generic customisation facilities (\secref{sec:program}). |
29 we discuss the generic customisation facilities (\secref{sec:program}). |
31 A further section (\secref{sec:adaptation}) is dedicated to the matter of |
30 A further section (\secref{sec:adaptation}) is dedicated to the matter of |
32 \qn{adaptation} to specific target language environments. After some |
31 \qn{adaptation} to specific target language environments. After some |
33 further issues (\secref{sec:further}) we conclude with an overview |
32 further issues (\secref{sec:further}) we conclude with an overview |
56 classes are identified with corresponding concepts in the target language. |
55 classes are identified with corresponding concepts in the target language. |
57 |
56 |
58 Inside @{theory HOL}, the @{command datatype} and |
57 Inside @{theory HOL}, the @{command datatype} and |
59 @{command definition}/@{command primrec}/@{command fun} declarations form |
58 @{command definition}/@{command primrec}/@{command fun} declarations form |
60 the core of a functional programming language. The default code generator setup |
59 the core of a functional programming language. The default code generator setup |
61 allows to turn those into functional programs immediately. |
60 transforms those into functional programs immediately. |
62 This means that \qt{naive} code generation can proceed without further ado. |
61 This means that \qt{naive} code generation can proceed without further ado. |
63 For example, here a simple \qt{implementation} of amortised queues: |
62 For example, here a simple \qt{implementation} of amortised queues: |
64 *} |
63 *} |
65 |
64 |
66 datatype %quote 'a queue = AQueue "'a list" "'a list" |
65 datatype %quote 'a queue = AQueue "'a list" "'a list" |
115 subsection {* Code generator architecture \label{sec:concept} *} |
114 subsection {* Code generator architecture \label{sec:concept} *} |
116 |
115 |
117 text {* |
116 text {* |
118 What you have seen so far should be already enough in a lot of cases. If you |
117 What you have seen so far should be already enough in a lot of cases. If you |
119 are content with this, you can quit reading here. Anyway, in order to customise |
118 are content with this, you can quit reading here. Anyway, in order to customise |
120 and adapt the code generator, it is inevitable to gain some understanding |
119 and adapt the code generator, it is necessary to gain some understanding |
121 how it works. |
120 how it works. |
122 |
121 |
123 \begin{figure}[h] |
122 \begin{figure}[h] |
124 \includegraphics{architecture} |
123 \includegraphics{architecture} |
125 \caption{Code generator architecture} |
124 \caption{Code generator architecture} |
140 the input |
139 the input |
141 of the next in the chain, see figure \ref{fig:arch}: |
140 of the next in the chain, see figure \ref{fig:arch}: |
142 |
141 |
143 \begin{itemize} |
142 \begin{itemize} |
144 |
143 |
145 \item Starting point is a collection of raw code equations in a |
144 \item The starting point is a collection of raw code equations in a |
146 theory; due to proof irrelevance it is not relevant where they |
145 theory. It is not relevant where they |
147 stem from but typically they are either descendant of specification |
146 stem from, but typically they were either produced by specification |
148 tools or explicit proofs by the user. |
147 tools or proved explicitly by the user. |
149 |
148 |
150 \item Before these raw code equations are continued |
149 \item These raw code equations can be subjected to theorem transformations. This |
151 with, they can be subjected to theorem transformations. This |
150 \qn{preprocessor} can apply the full |
152 \qn{preprocessor} is an interface which allows to apply the full |
|
153 expressiveness of ML-based theorem transformations to code |
151 expressiveness of ML-based theorem transformations to code |
154 generation. The result of the preprocessing step is a |
152 generation. The result of preprocessing is a |
155 structured collection of code equations. |
153 structured collection of code equations. |
156 |
154 |
157 \item These code equations are \qn{translated} to a program in an |
155 \item These code equations are \qn{translated} to a program in an |
158 abstract intermediate language. Think of it as a kind |
156 abstract intermediate language. Think of it as a kind |
159 of \qt{Mini-Haskell} with four \qn{statements}: @{text data} |
157 of \qt{Mini-Haskell} with four \qn{statements}: @{text data} |