| 28213 |      1 | theory Further
 | 
|  |      2 | imports Setup
 | 
|  |      3 | begin
 | 
|  |      4 | 
 | 
| 28447 |      5 | section {* Further issues \label{sec:further} *}
 | 
| 28419 |      6 | 
 | 
|  |      7 | subsection {* Further reading *}
 | 
|  |      8 | 
 | 
|  |      9 | text {*
 | 
|  |     10 |   Do dive deeper into the issue of code generation, you should visit
 | 
|  |     11 |   the Isabelle/Isar Reference Manual \cite{isabelle-isar-ref} which
 | 
| 28593 |     12 |   contains exhaustive syntax diagrams.
 | 
| 28419 |     13 | *}
 | 
|  |     14 | 
 | 
|  |     15 | subsection {* Modules *}
 | 
|  |     16 | 
 | 
|  |     17 | text {*
 | 
|  |     18 |   When invoking the @{command export_code} command it is possible to leave
 | 
|  |     19 |   out the @{keyword "module_name"} part;  then code is distributed over
 | 
|  |     20 |   different modules, where the module name space roughly is induced
 | 
| 28635 |     21 |   by the @{text Isabelle} theory name space.
 | 
| 28213 |     22 | 
 | 
| 28419 |     23 |   Then sometimes the awkward situation occurs that dependencies between
 | 
|  |     24 |   definitions introduce cyclic dependencies between modules, which in the
 | 
|  |     25 |   @{text Haskell} world leaves you to the mercy of the @{text Haskell} implementation
 | 
|  |     26 |   you are using,  while for @{text SML}/@{text OCaml} code generation is not possible.
 | 
|  |     27 | 
 | 
|  |     28 |   A solution is to declare module names explicitly.
 | 
|  |     29 |   Let use assume the three cyclically dependent
 | 
|  |     30 |   modules are named \emph{A}, \emph{B} and \emph{C}.
 | 
|  |     31 |   Then, by stating
 | 
|  |     32 | *}
 | 
|  |     33 | 
 | 
| 28593 |     34 | code_modulename %quote SML
 | 
| 28419 |     35 |   A ABC
 | 
|  |     36 |   B ABC
 | 
|  |     37 |   C ABC
 | 
|  |     38 | 
 | 
|  |     39 | text {*
 | 
|  |     40 |   we explicitly map all those modules on \emph{ABC},
 | 
|  |     41 |   resulting in an ad-hoc merge of this three modules
 | 
|  |     42 |   at serialisation time.
 | 
|  |     43 | *}
 | 
| 28213 |     44 | 
 | 
|  |     45 | subsection {* Evaluation oracle *}
 | 
|  |     46 | 
 | 
| 28593 |     47 | text {*
 | 
|  |     48 |   Code generation may also be used to \emph{evaluate} expressions
 | 
|  |     49 |   (using @{text SML} as target language of course).
 | 
|  |     50 |   For instance, the @{command value} allows to reduce an expression to a
 | 
|  |     51 |   normal form with respect to the underlying code equations:
 | 
|  |     52 | *}
 | 
|  |     53 | 
 | 
|  |     54 | value %quote "42 / (12 :: rat)"
 | 
|  |     55 | 
 | 
|  |     56 | text {*
 | 
|  |     57 |   \noindent will display @{term "7 / (2 :: rat)"}.
 | 
|  |     58 | 
 | 
|  |     59 |   The @{method eval} method tries to reduce a goal by code generation to @{term True}
 | 
|  |     60 |   and solves it in that case, but fails otherwise:
 | 
|  |     61 | *}
 | 
|  |     62 | 
 | 
|  |     63 | lemma %quote "42 / (12 :: rat) = 7 / 2"
 | 
|  |     64 |   by %quote eval
 | 
|  |     65 | 
 | 
|  |     66 | text {*
 | 
|  |     67 |   \noindent The soundness of the @{method eval} method depends crucially 
 | 
|  |     68 |   on the correctness of the code generator;  this is one of the reasons
 | 
| 31050 |     69 |   why you should not use adaptation (see \secref{sec:adaptation}) frivolously.
 | 
| 28593 |     70 | *}
 | 
|  |     71 | 
 | 
| 28213 |     72 | subsection {* Code antiquotation *}
 | 
|  |     73 | 
 | 
| 28593 |     74 | text {*
 | 
|  |     75 |   In scenarios involving techniques like reflection it is quite common
 | 
|  |     76 |   that code generated from a theory forms the basis for implementing
 | 
|  |     77 |   a proof procedure in @{text SML}.  To facilitate interfacing of generated code
 | 
|  |     78 |   with system code, the code generator provides a @{text code} antiquotation:
 | 
|  |     79 | *}
 | 
|  |     80 | 
 | 
| 30880 |     81 | datatype %quote form = T | F | And form form | Or form form (*<*)
 | 
|  |     82 | 
 | 
|  |     83 | (*>*) ML %quotett {*
 | 
| 28593 |     84 |   fun eval_form @{code T} = true
 | 
|  |     85 |     | eval_form @{code F} = false
 | 
|  |     86 |     | eval_form (@{code And} (p, q)) =
 | 
|  |     87 |         eval_form p andalso eval_form q
 | 
|  |     88 |     | eval_form (@{code Or} (p, q)) =
 | 
|  |     89 |         eval_form p orelse eval_form q;
 | 
|  |     90 | *}
 | 
|  |     91 | 
 | 
|  |     92 | text {*
 | 
|  |     93 |   \noindent @{text code} takes as argument the name of a constant;  after the
 | 
|  |     94 |   whole @{text SML} is read, the necessary code is generated transparently
 | 
|  |     95 |   and the corresponding constant names are inserted.  This technique also
 | 
|  |     96 |   allows to use pattern matching on constructors stemming from compiled
 | 
|  |     97 |   @{text datatypes}.
 | 
|  |     98 | 
 | 
| 29796 |     99 |   For a less simplistic example, theory @{theory Ferrack} is
 | 
| 28593 |    100 |   a good reference.
 | 
|  |    101 | *}
 | 
| 28213 |    102 | 
 | 
| 28419 |    103 | subsection {* Imperative data structures *}
 | 
|  |    104 | 
 | 
| 28593 |    105 | text {*
 | 
|  |    106 |   If you consider imperative data structures as inevitable for a specific
 | 
|  |    107 |   application, you should consider
 | 
|  |    108 |   \emph{Imperative Functional Programming with Isabelle/HOL}
 | 
|  |    109 |   (\cite{bulwahn-et-al:2008:imperative});
 | 
|  |    110 |   the framework described there is available in theory @{theory Imperative_HOL}.
 | 
|  |    111 | *}
 | 
|  |    112 | 
 | 
| 28213 |    113 | end
 |