| author | haftmann | 
| Mon, 31 May 2010 17:29:28 +0200 | |
| changeset 37211 | 32a6f471f090 | 
| parent 34179 | 5490151d1052 | 
| child 37428 | b3d94253e7f2 | 
| permissions | -rw-r--r-- | 
| 28447 | 1 | % | 
| 2 | \begin{isabellebody}%
 | |
| 3 | \def\isabellecontext{Introduction}%
 | |
| 4 | % | |
| 5 | \isadelimtheory | |
| 6 | % | |
| 7 | \endisadelimtheory | |
| 8 | % | |
| 9 | \isatagtheory | |
| 10 | \isacommand{theory}\isamarkupfalse%
 | |
| 11 | \ Introduction\isanewline | |
| 12 | \isakeyword{imports}\ Setup\isanewline
 | |
| 13 | \isakeyword{begin}%
 | |
| 14 | \endisatagtheory | |
| 15 | {\isafoldtheory}%
 | |
| 16 | % | |
| 17 | \isadelimtheory | |
| 18 | % | |
| 19 | \endisadelimtheory | |
| 20 | % | |
| 21 | \isamarkupsection{Introduction and Overview%
 | |
| 22 | } | |
| 23 | \isamarkuptrue% | |
| 24 | % | |
| 25 | \begin{isamarkuptext}%
 | |
| 26 | This tutorial introduces a generic code generator for the | |
| 27 |   \isa{Isabelle} system.
 | |
| 34155 | 28 | The | 
| 29 |   \qn{target language} for which code is
 | |
| 30 | generated is not fixed, but may be one of several | |
| 31 | functional programming languages (currently, the implementation | |
| 28447 | 32 |   supports \isa{SML} \cite{SML}, \isa{OCaml} \cite{OCaml} and \isa{Haskell}
 | 
| 33 |   \cite{haskell-revised-report}).
 | |
| 34 | ||
| 35 | Conceptually the code generator framework is part | |
| 36 |   of Isabelle's \hyperlink{theory.Pure}{\mbox{\isa{Pure}}} meta logic framework; the logic
 | |
| 34155 | 37 |   \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} \cite{isa-tutorial}, which is an extension of \hyperlink{theory.Pure}{\mbox{\isa{Pure}}},
 | 
| 28447 | 38 | already comes with a reasonable framework setup and thus provides | 
| 34155 | 39 | a good basis for creating code-generation-driven | 
| 28447 | 40 | applications. So, we assume some familiarity and experience | 
| 41 |   with the ingredients of the \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} distribution theories.
 | |
| 42 | ||
| 43 | The code generator aims to be usable with no further ado | |
| 34155 | 44 | in most cases, while allowing for detailed customisation. | 
| 45 | This can be seen in the structure of this tutorial: after a short | |
| 28447 | 46 |   conceptual introduction with an example (\secref{sec:intro}),
 | 
| 47 |   we discuss the generic customisation facilities (\secref{sec:program}).
 | |
| 31050 | 48 |   A further section (\secref{sec:adaptation}) is dedicated to the matter of
 | 
| 49 |   \qn{adaptation} to specific target language environments.  After some
 | |
| 28447 | 50 |   further issues (\secref{sec:further}) we conclude with an overview
 | 
| 51 |   of some ML programming interfaces (\secref{sec:ml}).
 | |
| 52 | ||
| 53 |   \begin{warn}
 | |
| 54 | Ultimately, the code generator which this tutorial deals with | |
| 55 | is supposed to replace the existing code generator | |
| 56 |     by Stefan Berghofer \cite{Berghofer-Nipkow:2002}.
 | |
| 57 | So, for the moment, there are two distinct code generators | |
| 58 | in Isabelle. In case of ambiguity, we will refer to the framework | |
| 59 |     described here as \isa{generic\ code\ generator}, to the
 | |
| 60 |     other as \isa{SML\ code\ generator}.
 | |
| 61 | Also note that while the framework itself is | |
| 62 |     object-logic independent, only \hyperlink{theory.HOL}{\mbox{\isa{HOL}}} provides a reasonable
 | |
| 63 | framework setup. | |
| 64 |   \end{warn}%
 | |
| 65 | \end{isamarkuptext}%
 | |
| 66 | \isamarkuptrue% | |
| 67 | % | |
| 68 | \isamarkupsubsection{Code generation via shallow embedding \label{sec:intro}%
 | |
| 69 | } | |
| 70 | \isamarkuptrue% | |
| 71 | % | |
| 72 | \begin{isamarkuptext}%
 | |
| 73 | The key concept for understanding \isa{Isabelle}'s code generation is
 | |
| 74 |   \emph{shallow embedding}, i.e.~logical entities like constants, types and
 | |
| 75 | classes are identified with corresponding concepts in the target language. | |
| 76 | ||
| 77 |   Inside \hyperlink{theory.HOL}{\mbox{\isa{HOL}}}, the \hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}} and
 | |
| 78 |   \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}/\hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}}/\hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}} declarations form
 | |
| 79 | the core of a functional programming language. The default code generator setup | |
| 34155 | 80 | transforms those into functional programs immediately. | 
| 28447 | 81 |   This means that \qt{naive} code generation can proceed without further ado.
 | 
| 82 |   For example, here a simple \qt{implementation} of amortised queues:%
 | |
| 83 | \end{isamarkuptext}%
 | |
| 84 | \isamarkuptrue% | |
| 85 | % | |
| 28564 | 86 | \isadelimquote | 
| 28447 | 87 | % | 
| 28564 | 88 | \endisadelimquote | 
| 28447 | 89 | % | 
| 28564 | 90 | \isatagquote | 
| 28447 | 91 | \isacommand{datatype}\isamarkupfalse%
 | 
| 29798 | 92 | \ {\isacharprime}a\ queue\ {\isacharequal}\ AQueue\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharprime}a\ list{\isachardoublequoteclose}\isanewline
 | 
| 28447 | 93 | \isanewline | 
| 94 | \isacommand{definition}\isamarkupfalse%
 | |
| 95 | \ empty\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | |
| 29798 | 96 | \ \ {\isachardoublequoteopen}empty\ {\isacharequal}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
 | 
| 28447 | 97 | \isanewline | 
| 98 | \isacommand{primrec}\isamarkupfalse%
 | |
| 99 | \ enqueue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | |
| 29798 | 100 | \ \ {\isachardoublequoteopen}enqueue\ x\ {\isacharparenleft}AQueue\ xs\ ys{\isacharparenright}\ {\isacharequal}\ AQueue\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ ys{\isachardoublequoteclose}\isanewline
 | 
| 28447 | 101 | \isanewline | 
| 102 | \isacommand{fun}\isamarkupfalse%
 | |
| 103 | \ dequeue\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ queue\ {\isasymRightarrow}\ {\isacharprime}a\ option\ {\isasymtimes}\ {\isacharprime}a\ queue{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 | |
| 29798 | 104 | \ \ \ \ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}None{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | 
| 105 | \ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharparenleft}y\ {\isacharhash}\ ys{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ xs\ ys{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 106 | \ \ {\isacharbar}\ {\isachardoublequoteopen}dequeue\ {\isacharparenleft}AQueue\ xs\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\isanewline
 | |
| 107 | \ \ \ \ \ \ {\isacharparenleft}case\ rev\ xs\ of\ y\ {\isacharhash}\ ys\ {\isasymRightarrow}\ {\isacharparenleft}Some\ y{\isacharcomma}\ AQueue\ {\isacharbrackleft}{\isacharbrackright}\ ys{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}%
 | |
| 28564 | 108 | \endisatagquote | 
| 109 | {\isafoldquote}%
 | |
| 28447 | 110 | % | 
| 28564 | 111 | \isadelimquote | 
| 28447 | 112 | % | 
| 28564 | 113 | \endisadelimquote | 
| 28447 | 114 | % | 
| 115 | \begin{isamarkuptext}%
 | |
| 116 | \noindent Then we can generate code e.g.~for \isa{SML} as follows:%
 | |
| 117 | \end{isamarkuptext}%
 | |
| 118 | \isamarkuptrue% | |
| 119 | % | |
| 28564 | 120 | \isadelimquote | 
| 28447 | 121 | % | 
| 28564 | 122 | \endisadelimquote | 
| 28447 | 123 | % | 
| 28564 | 124 | \isatagquote | 
| 28447 | 125 | \isacommand{export{\isacharunderscore}code}\isamarkupfalse%
 | 
| 126 | \ empty\ dequeue\ enqueue\ \isakeyword{in}\ SML\isanewline
 | |
| 127 | \ \ \isakeyword{module{\isacharunderscore}name}\ Example\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}example{\isachardot}ML{\isachardoublequoteclose}%
 | |
| 28564 | 128 | \endisatagquote | 
| 129 | {\isafoldquote}%
 | |
| 28447 | 130 | % | 
| 28564 | 131 | \isadelimquote | 
| 28447 | 132 | % | 
| 28564 | 133 | \endisadelimquote | 
| 28447 | 134 | % | 
| 135 | \begin{isamarkuptext}%
 | |
| 136 | \noindent resulting in the following code:% | |
| 137 | \end{isamarkuptext}%
 | |
| 138 | \isamarkuptrue% | |
| 139 | % | |
| 28564 | 140 | \isadelimquote | 
| 28447 | 141 | % | 
| 28564 | 142 | \endisadelimquote | 
| 28447 | 143 | % | 
| 28564 | 144 | \isatagquote | 
| 28447 | 145 | % | 
| 146 | \begin{isamarkuptext}%
 | |
| 28727 | 147 | \isatypewriter% | 
| 28447 | 148 | \noindent% | 
| 34155 | 149 | \hspace*{0pt}structure Example :~sig\\
 | 
| 34179 | 150 | \hspace*{0pt} ~val foldl :~('a -> 'b -> 'a) -> 'a -> 'b list -> 'a\\
 | 
| 34155 | 151 | \hspace*{0pt} ~val rev :~'a list -> 'a list\\
 | 
| 34179 | 152 | \hspace*{0pt} ~val list{\char95}case :~'a -> ('b -> 'b list -> 'a) -> 'b list -> 'a\\
 | 
| 153 | \hspace*{0pt} ~datatype 'a queue = AQueue of 'a list * 'a list\\
 | |
| 34155 | 154 | \hspace*{0pt} ~val empty :~'a queue\\
 | 
| 155 | \hspace*{0pt} ~val dequeue :~'a queue -> 'a option * 'a queue\\
 | |
| 156 | \hspace*{0pt} ~val enqueue :~'a -> 'a queue -> 'a queue\\
 | |
| 157 | \hspace*{0pt}end = struct\\
 | |
| 28714 | 158 | \hspace*{0pt}\\
 | 
| 159 | \hspace*{0pt}fun foldl f a [] = a\\
 | |
| 34179 | 160 | \hspace*{0pt} ~| foldl f a (x ::~xs) = foldl f (f a x) xs;\\
 | 
| 28714 | 161 | \hspace*{0pt}\\
 | 
| 34179 | 162 | \hspace*{0pt}fun rev xs = foldl (fn xsa => fn x => x ::~xsa) [] xs;\\
 | 
| 28714 | 163 | \hspace*{0pt}\\
 | 
| 34179 | 164 | \hspace*{0pt}fun list{\char95}case f1 f2 (a ::~lista) = f2 a lista\\
 | 
| 165 | \hspace*{0pt} ~| list{\char95}case f1 f2 [] = f1;\\
 | |
| 28714 | 166 | \hspace*{0pt}\\
 | 
| 34179 | 167 | \hspace*{0pt}datatype 'a queue = AQueue of 'a list * 'a list;\\
 | 
| 28714 | 168 | \hspace*{0pt}\\
 | 
| 34179 | 169 | \hspace*{0pt}val empty :~'a queue = AQueue ([],~[]);\\
 | 
| 28714 | 170 | \hspace*{0pt}\\
 | 
| 34179 | 171 | \hspace*{0pt}fun dequeue (AQueue ([],~[])) = (NONE,~AQueue ([],~[]))\\
 | 
| 172 | \hspace*{0pt} ~| dequeue (AQueue (xs,~y ::~ys)) = (SOME y,~AQueue (xs,~ys))\\
 | |
| 173 | \hspace*{0pt} ~| dequeue (AQueue (v ::~va,~[])) =\\
 | |
| 28714 | 174 | \hspace*{0pt} ~~~let\\
 | 
| 34179 | 175 | \hspace*{0pt} ~~~~~val y ::~ys = rev (v ::~va);\\
 | 
| 28714 | 176 | \hspace*{0pt} ~~~in\\
 | 
| 34179 | 177 | \hspace*{0pt} ~~~~~(SOME y,~AQueue ([],~ys))\\
 | 
| 28714 | 178 | \hspace*{0pt} ~~~end;\\
 | 
| 179 | \hspace*{0pt}\\
 | |
| 34179 | 180 | \hspace*{0pt}fun enqueue x (AQueue (xs,~ys)) = AQueue (x ::~xs,~ys);\\
 | 
| 28714 | 181 | \hspace*{0pt}\\
 | 
| 29297 | 182 | \hspace*{0pt}end;~(*struct Example*)%
 | 
| 28447 | 183 | \end{isamarkuptext}%
 | 
| 184 | \isamarkuptrue% | |
| 185 | % | |
| 28564 | 186 | \endisatagquote | 
| 187 | {\isafoldquote}%
 | |
| 28447 | 188 | % | 
| 28564 | 189 | \isadelimquote | 
| 28447 | 190 | % | 
| 28564 | 191 | \endisadelimquote | 
| 28447 | 192 | % | 
| 193 | \begin{isamarkuptext}%
 | |
| 194 | \noindent The \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command takes a space-separated list of
 | |
| 195 | constants for which code shall be generated; anything else needed for those | |
| 196 | is added implicitly. Then follows a target language identifier | |
| 197 |   (\isa{SML}, \isa{OCaml} or \isa{Haskell}) and a freely chosen module name.
 | |
| 198 | A file name denotes the destination to store the generated code. Note that | |
| 199 | the semantics of the destination depends on the target language: for | |
| 200 |   \isa{SML} and \isa{OCaml} it denotes a \emph{file}, for \isa{Haskell}
 | |
| 201 |   it denotes a \emph{directory} where a file named as the module name
 | |
| 202 |   (with extension \isa{{\isachardot}hs}) is written:%
 | |
| 203 | \end{isamarkuptext}%
 | |
| 204 | \isamarkuptrue% | |
| 205 | % | |
| 28564 | 206 | \isadelimquote | 
| 28447 | 207 | % | 
| 28564 | 208 | \endisadelimquote | 
| 28447 | 209 | % | 
| 28564 | 210 | \isatagquote | 
| 28447 | 211 | \isacommand{export{\isacharunderscore}code}\isamarkupfalse%
 | 
| 212 | \ empty\ dequeue\ enqueue\ \isakeyword{in}\ Haskell\isanewline
 | |
| 213 | \ \ \isakeyword{module{\isacharunderscore}name}\ Example\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}{\isachardoublequoteclose}%
 | |
| 28564 | 214 | \endisatagquote | 
| 215 | {\isafoldquote}%
 | |
| 28447 | 216 | % | 
| 28564 | 217 | \isadelimquote | 
| 28447 | 218 | % | 
| 28564 | 219 | \endisadelimquote | 
| 28447 | 220 | % | 
| 221 | \begin{isamarkuptext}%
 | |
| 34155 | 222 | \noindent This is the corresponding code in \isa{Haskell}:%
 | 
| 28447 | 223 | \end{isamarkuptext}%
 | 
| 224 | \isamarkuptrue% | |
| 225 | % | |
| 28564 | 226 | \isadelimquote | 
| 28447 | 227 | % | 
| 28564 | 228 | \endisadelimquote | 
| 28447 | 229 | % | 
| 28564 | 230 | \isatagquote | 
| 28447 | 231 | % | 
| 232 | \begin{isamarkuptext}%
 | |
| 28727 | 233 | \isatypewriter% | 
| 28447 | 234 | \noindent% | 
| 28714 | 235 | \hspace*{0pt}module Example where {\char123}\\
 | 
| 236 | \hspace*{0pt}\\
 | |
| 34179 | 237 | \hspace*{0pt}foldla ::~forall a b.~(a -> b -> a) -> a -> [b] -> a;\\
 | 
| 28714 | 238 | \hspace*{0pt}foldla f a [] = a;\\
 | 
| 34179 | 239 | \hspace*{0pt}foldla f a (x :~xs) = foldla f (f a x) xs;\\
 | 
| 28714 | 240 | \hspace*{0pt}\\
 | 
| 29297 | 241 | \hspace*{0pt}rev ::~forall a.~[a] -> [a];\\
 | 
| 34179 | 242 | \hspace*{0pt}rev xs = foldla ({\char92}~xsa x -> x :~xsa) [] xs;\\
 | 
| 28714 | 243 | \hspace*{0pt}\\
 | 
| 34179 | 244 | \hspace*{0pt}list{\char95}case ::~forall a b.~a -> (b -> [b] -> a) -> [b] -> a;\\
 | 
| 245 | \hspace*{0pt}list{\char95}case f1 f2 (a :~list) = f2 a list;\\
 | |
| 28714 | 246 | \hspace*{0pt}list{\char95}case f1 f2 [] = f1;\\
 | 
| 247 | \hspace*{0pt}\\
 | |
| 29798 | 248 | \hspace*{0pt}data Queue a = AQueue [a] [a];\\
 | 
| 28714 | 249 | \hspace*{0pt}\\
 | 
| 29297 | 250 | \hspace*{0pt}empty ::~forall a.~Queue a;\\
 | 
| 29798 | 251 | \hspace*{0pt}empty = AQueue [] [];\\
 | 
| 28714 | 252 | \hspace*{0pt}\\
 | 
| 29297 | 253 | \hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\
 | 
| 34179 | 254 | \hspace*{0pt}dequeue (AQueue [] []) = (Nothing,~AQueue [] []);\\
 | 
| 255 | \hspace*{0pt}dequeue (AQueue xs (y :~ys)) = (Just y,~AQueue xs ys);\\
 | |
| 256 | \hspace*{0pt}dequeue (AQueue (v :~va) []) =\\
 | |
| 31848 | 257 | \hspace*{0pt} ~let {\char123}\\
 | 
| 34179 | 258 | \hspace*{0pt} ~~~(y :~ys) = rev (v :~va);\\
 | 
| 31848 | 259 | \hspace*{0pt} ~{\char125}~in (Just y,~AQueue [] ys);\\
 | 
| 28714 | 260 | \hspace*{0pt}\\
 | 
| 29297 | 261 | \hspace*{0pt}enqueue ::~forall a.~a -> Queue a -> Queue a;\\
 | 
| 34179 | 262 | \hspace*{0pt}enqueue x (AQueue xs ys) = AQueue (x :~xs) ys;\\
 | 
| 28714 | 263 | \hspace*{0pt}\\
 | 
| 264 | \hspace*{0pt}{\char125}%
 | |
| 28447 | 265 | \end{isamarkuptext}%
 | 
| 266 | \isamarkuptrue% | |
| 267 | % | |
| 28564 | 268 | \endisatagquote | 
| 269 | {\isafoldquote}%
 | |
| 28447 | 270 | % | 
| 28564 | 271 | \isadelimquote | 
| 28447 | 272 | % | 
| 28564 | 273 | \endisadelimquote | 
| 28447 | 274 | % | 
| 275 | \begin{isamarkuptext}%
 | |
| 276 | \noindent This demonstrates the basic usage of the \hyperlink{command.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} command;
 | |
| 277 |   for more details see \secref{sec:further}.%
 | |
| 278 | \end{isamarkuptext}%
 | |
| 279 | \isamarkuptrue% | |
| 280 | % | |
| 281 | \isamarkupsubsection{Code generator architecture \label{sec:concept}%
 | |
| 282 | } | |
| 283 | \isamarkuptrue% | |
| 284 | % | |
| 285 | \begin{isamarkuptext}%
 | |
| 286 | What you have seen so far should be already enough in a lot of cases. If you | |
| 287 | are content with this, you can quit reading here. Anyway, in order to customise | |
| 34155 | 288 | and adapt the code generator, it is necessary to gain some understanding | 
| 28447 | 289 | how it works. | 
| 290 | ||
| 291 |   \begin{figure}[h]
 | |
| 30882 
d15725e84091
moved generated eps/pdf to main directory, for proper display in dvi;
 wenzelm parents: 
30880diff
changeset | 292 |     \includegraphics{architecture}
 | 
| 28447 | 293 |     \caption{Code generator architecture}
 | 
| 294 |     \label{fig:arch}
 | |
| 295 |   \end{figure}
 | |
| 296 | ||
| 297 | The code generator employs a notion of executability | |
| 298 | for three foundational executable ingredients known | |
| 299 | from functional programming: | |
| 29560 | 300 |   \emph{code equations}, \emph{datatypes}, and
 | 
| 301 |   \emph{type classes}.  A code equation as a first approximation
 | |
| 28447 | 302 |   is a theorem of the form \isa{f\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n\ {\isasymequiv}\ t}
 | 
| 303 |   (an equation headed by a constant \isa{f} with arguments
 | |
| 304 |   \isa{t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n} and right hand side \isa{t}).
 | |
| 29560 | 305 | Code generation aims to turn code equations | 
| 28447 | 306 | into a functional program. This is achieved by three major | 
| 307 | components which operate sequentially, i.e. the result of one is | |
| 308 | the input | |
| 30880 | 309 |   of the next in the chain,  see figure \ref{fig:arch}:
 | 
| 28447 | 310 | |
| 311 |   \begin{itemize}
 | |
| 312 | ||
| 34155 | 313 | \item The starting point is a collection of raw code equations in a | 
| 314 | theory. It is not relevant where they | |
| 315 | stem from, but typically they were either produced by specification | |
| 316 | tools or proved explicitly by the user. | |
| 30836 | 317 | |
| 34155 | 318 | \item These raw code equations can be subjected to theorem transformations. This | 
| 319 |       \qn{preprocessor} can apply the full
 | |
| 30836 | 320 | expressiveness of ML-based theorem transformations to code | 
| 34155 | 321 | generation. The result of preprocessing is a | 
| 30836 | 322 | structured collection of code equations. | 
| 28447 | 323 | |
| 30836 | 324 |     \item These code equations are \qn{translated} to a program in an
 | 
| 325 | abstract intermediate language. Think of it as a kind | |
| 28447 | 326 |       of \qt{Mini-Haskell} with four \qn{statements}: \isa{data}
 | 
| 29560 | 327 |       (for datatypes), \isa{fun} (stemming from code equations),
 | 
| 28447 | 328 |       also \isa{class} and \isa{inst} (for type classes).
 | 
| 329 | ||
| 330 |     \item Finally, the abstract program is \qn{serialised} into concrete
 | |
| 331 | source code of a target language. | |
| 30836 | 332 | This step only produces concrete syntax but does not change the | 
| 333 | program in essence; all conceptual transformations occur in the | |
| 334 | translation step. | |
| 28447 | 335 | |
| 336 |   \end{itemize}
 | |
| 337 | ||
| 338 | \noindent From these steps, only the two last are carried out outside the logic; by | |
| 339 | keeping this layer as thin as possible, the amount of code to trust is | |
| 340 | kept to a minimum.% | |
| 341 | \end{isamarkuptext}%
 | |
| 342 | \isamarkuptrue% | |
| 343 | % | |
| 344 | \isadelimtheory | |
| 345 | % | |
| 346 | \endisadelimtheory | |
| 347 | % | |
| 348 | \isatagtheory | |
| 349 | \isacommand{end}\isamarkupfalse%
 | |
| 350 | % | |
| 351 | \endisatagtheory | |
| 352 | {\isafoldtheory}%
 | |
| 353 | % | |
| 354 | \isadelimtheory | |
| 355 | % | |
| 356 | \endisadelimtheory | |
| 357 | \isanewline | |
| 358 | \end{isabellebody}%
 | |
| 359 | %%% Local Variables: | |
| 360 | %%% mode: latex | |
| 361 | %%% TeX-master: "root" | |
| 362 | %%% End: |