| author | wenzelm | 
| Thu, 02 Mar 2023 16:09:22 +0100 | |
| changeset 77483 | 291f5848bf55 | 
| parent 76987 | 4c275405faae | 
| child 78665 | b0ddfa5b9ddc | 
| permissions | -rw-r--r-- | 
| 28213 | 1 | theory Further | 
| 69422 
472af2d7835d
clarified session dependencies: faster build_doc/build_release;
 wenzelm parents: 
68484diff
changeset | 2 | imports Setup | 
| 28213 | 3 | begin | 
| 4 | ||
| 59377 | 5 | section \<open>Further issues \label{sec:further}\<close>
 | 
| 28419 | 6 | |
| 69505 | 7 | subsection \<open>Incorporating generated code directly into the system runtime -- \<open>code_reflect\<close>\<close> | 
| 65041 | 8 | |
| 9 | subsubsection \<open>Static embedding of generated code into the system runtime\<close> | |
| 10 | ||
| 11 | text \<open> | |
| 12 |   The @{ML_antiquotation code} antiquotation is lightweight, but the generated code
 | |
| 13 | is only accessible while the ML section is processed. Sometimes this | |
| 14 | is not appropriate, especially if the generated code contains datatype | |
| 15 | declarations which are shared with other parts of the system. In these | |
| 16 |   cases, @{command_def code_reflect} can be used:
 | |
| 17 | \<close> | |
| 18 | ||
| 19 | code_reflect %quote Sum_Type | |
| 20 | datatypes sum = Inl | Inr | |
| 21 | functions "Sum_Type.sum.projl" "Sum_Type.sum.projr" | |
| 22 | ||
| 23 | text \<open> | |
| 24 |   \noindent @{command code_reflect} takes a structure name and
 | |
| 25 | references to datatypes and functions; for these code is compiled | |
| 26 |   into the named ML structure and the \emph{Eval} target is modified
 | |
| 27 | in a way that future code generation will reference these | |
| 28 | precompiled versions of the given datatypes and functions. This | |
| 29 | also allows to refer to the referenced datatypes and functions from | |
| 30 | arbitrary ML code as well. | |
| 31 | ||
| 32 |   A typical example for @{command code_reflect} can be found in the
 | |
| 69597 | 33 | \<^theory>\<open>HOL.Predicate\<close> theory. | 
| 65041 | 34 | \<close> | 
| 35 | ||
| 36 | ||
| 37 | subsubsection \<open>Separate compilation\<close> | |
| 38 | ||
| 39 | text \<open> | |
| 40 | For technical reasons it is sometimes necessary to separate | |
| 41 | generation and compilation of code which is supposed to be used in | |
| 42 |   the system runtime.  For this @{command code_reflect} with an
 | |
| 75074 
78c2a92a8be4
updated documentation to current matter of affairs
 haftmann parents: 
70022diff
changeset | 43 | optional \<^theory_text>\<open>file_prefix\<close> argument can be used: | 
| 65041 | 44 | \<close> | 
| 45 | ||
| 46 | code_reflect %quote Rat | |
| 47 | datatypes rat | |
| 48 | functions Fract | |
| 49 | "(plus :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(minus :: rat \<Rightarrow> rat \<Rightarrow> rat)" | |
| 50 | "(times :: rat \<Rightarrow> rat \<Rightarrow> rat)" "(divide :: rat \<Rightarrow> rat \<Rightarrow> rat)" | |
| 70022 
49e178cbf923
'code_reflect' only supports new-style 'file_prefix';
 wenzelm parents: 
69697diff
changeset | 51 | file_prefix rat | 
| 65041 | 52 | |
| 53 | text \<open> | |
| 75074 
78c2a92a8be4
updated documentation to current matter of affairs
 haftmann parents: 
70022diff
changeset | 54 | \noindent This generates the referenced code as logical files within the theory context, | 
| 
78c2a92a8be4
updated documentation to current matter of affairs
 haftmann parents: 
70022diff
changeset | 55 |   similar to @{command export_code}.
 | 
| 65041 | 56 | \<close> | 
| 57 | ||
| 69505 | 58 | subsection \<open>Specialities of the \<open>Scala\<close> target language \label{sec:scala}\<close>
 | 
| 42096 | 59 | |
| 59377 | 60 | text \<open> | 
| 69505 | 61 | \<open>Scala\<close> deviates from languages of the ML family in a couple | 
| 42096 | 62 | of aspects; those which affect code generation mainly have to do with | 
| 69505 | 63 | \<open>Scala\<close>'s type system: | 
| 42096 | 64 | |
| 65 |   \begin{itemize}
 | |
| 66 | ||
| 69505 | 67 | \item \<open>Scala\<close> prefers tupled syntax over curried syntax. | 
| 42096 | 68 | |
| 69505 | 69 | \item \<open>Scala\<close> sacrifices Hindely-Milner type inference for a | 
| 42096 | 70 | much more rich type system with subtyping etc. For this reason | 
| 71 | type arguments sometimes have to be given explicitly in square | |
| 46520 | 72 | brackets (mimicking System F syntax). | 
| 42096 | 73 | |
| 69505 | 74 | \item In contrast to \<open>Haskell\<close> where most specialities of | 
| 42096 | 75 |       the type system are implemented using \emph{type classes},
 | 
| 69505 | 76 |       \<open>Scala\<close> provides a sophisticated system of \emph{implicit
 | 
| 42096 | 77 | arguments}. | 
| 78 | ||
| 79 |   \end{itemize}
 | |
| 80 | ||
| 69505 | 81 | \noindent Concerning currying, the \<open>Scala\<close> serializer counts | 
| 42096 | 82 | arguments in code equations to determine how many arguments | 
| 83 | shall be tupled; remaining arguments and abstractions in terms | |
| 84 | rather than function definitions are always curried. | |
| 85 | ||
| 86 |   The second aspect affects user-defined adaptations with @{command
 | |
| 69505 | 87 | code_printing}. For regular terms, the \<open>Scala\<close> serializer prints | 
| 42096 | 88 | all type arguments explicitly. For user-defined term adaptations | 
| 89 | this is only possible for adaptations which take no arguments: here | |
| 90 | the type arguments are just appended. Otherwise they are ignored; | |
| 91 | hence user-defined adaptations for polymorphic constants have to be | |
| 92 | designed very carefully to avoid ambiguity. | |
| 67207 
ad538f6c5d2f
dedicated case option for code generation to Scala
 haftmann parents: 
66453diff
changeset | 93 | |
| 
ad538f6c5d2f
dedicated case option for code generation to Scala
 haftmann parents: 
66453diff
changeset | 94 | Note also that name clashes can occur when generating Scala code | 
| 
ad538f6c5d2f
dedicated case option for code generation to Scala
 haftmann parents: 
66453diff
changeset | 95 | to case-insensitive file systems; these can be avoid using the | 
| 
ad538f6c5d2f
dedicated case option for code generation to Scala
 haftmann parents: 
66453diff
changeset | 96 |   ``\<open>(case_insensitive)\<close>'' option for @{command export_code}.
 | 
| 59377 | 97 | \<close> | 
| 42096 | 98 | |
| 99 | ||
| 59377 | 100 | subsection \<open>Modules namespace\<close> | 
| 28419 | 101 | |
| 59377 | 102 | text \<open> | 
| 40752 | 103 |   When invoking the @{command export_code} command it is possible to
 | 
| 104 |   leave out the @{keyword "module_name"} part; then code is
 | |
| 105 | distributed over different modules, where the module name space | |
| 106 | roughly is induced by the Isabelle theory name space. | |
| 38437 | 107 | |
| 40752 | 108 | Then sometimes the awkward situation occurs that dependencies | 
| 109 | between definitions introduce cyclic dependencies between modules, | |
| 69505 | 110 | which in the \<open>Haskell\<close> world leaves you to the mercy of the | 
| 111 | \<open>Haskell\<close> implementation you are using, while for \<open>SML\<close>/\<open>OCaml\<close> code generation is not possible. | |
| 38437 | 112 | |
| 40752 | 113 | A solution is to declare module names explicitly. Let use assume | 
| 114 |   the three cyclically dependent modules are named \emph{A}, \emph{B}
 | |
| 115 |   and \emph{C}.  Then, by stating
 | |
| 59377 | 116 | \<close> | 
| 38437 | 117 | |
| 52378 
08dbf9ff2140
documentation on code_printing and code_identifier
 haftmann parents: 
51717diff
changeset | 118 | code_identifier %quote | 
| 
08dbf9ff2140
documentation on code_printing and code_identifier
 haftmann parents: 
51717diff
changeset | 119 | code_module A \<rightharpoonup> (SML) ABC | 
| 
08dbf9ff2140
documentation on code_printing and code_identifier
 haftmann parents: 
51717diff
changeset | 120 | | code_module B \<rightharpoonup> (SML) ABC | 
| 
08dbf9ff2140
documentation on code_printing and code_identifier
 haftmann parents: 
51717diff
changeset | 121 | | code_module C \<rightharpoonup> (SML) ABC | 
| 38437 | 122 | |
| 59377 | 123 | text \<open> | 
| 40752 | 124 |   \noindent we explicitly map all those modules on \emph{ABC},
 | 
| 125 | resulting in an ad-hoc merge of this three modules at serialisation | |
| 126 | time. | |
| 59377 | 127 | \<close> | 
| 28419 | 128 | |
| 59377 | 129 | subsection \<open>Locales and interpretation\<close> | 
| 37426 | 130 | |
| 59377 | 131 | text \<open> | 
| 37426 | 132 | A technical issue comes to surface when generating code from | 
| 61891 
76189756ff65
documentation on last state of the art concerning interpretation
 haftmann parents: 
61779diff
changeset | 133 | specifications stemming from locale interpretation into global | 
| 
76189756ff65
documentation on last state of the art concerning interpretation
 haftmann parents: 
61779diff
changeset | 134 | theories. | 
| 37426 | 135 | |
| 40752 | 136 | Let us assume a locale specifying a power operation on arbitrary | 
| 137 | types: | |
| 59377 | 138 | \<close> | 
| 37426 | 139 | |
| 140 | locale %quote power = | |
| 141 | fixes power :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" | |
| 142 | assumes power_commute: "power x \<circ> power y = power y \<circ> power x" | |
| 143 | begin | |
| 144 | ||
| 59377 | 145 | text \<open> | 
| 69505 | 146 | \noindent Inside that locale we can lift \<open>power\<close> to exponent | 
| 61779 | 147 | lists by means of a specification relative to that locale: | 
| 59377 | 148 | \<close> | 
| 37426 | 149 | |
| 150 | primrec %quote powers :: "'a list \<Rightarrow> 'b \<Rightarrow> 'b" where | |
| 151 | "powers [] = id" | |
| 152 | | "powers (x # xs) = power x \<circ> powers xs" | |
| 153 | ||
| 154 | lemma %quote powers_append: | |
| 155 | "powers (xs @ ys) = powers xs \<circ> powers ys" | |
| 156 | by (induct xs) simp_all | |
| 157 | ||
| 158 | lemma %quote powers_power: | |
| 159 | "powers xs \<circ> power x = power x \<circ> powers xs" | |
| 160 | by (induct xs) | |
| 49739 | 161 | (simp_all del: o_apply id_apply add: comp_assoc, | 
| 37426 | 162 | simp del: o_apply add: o_assoc power_commute) | 
| 163 | ||
| 164 | lemma %quote powers_rev: | |
| 165 | "powers (rev xs) = powers xs" | |
| 166 | by (induct xs) (simp_all add: powers_append powers_power) | |
| 167 | ||
| 168 | end %quote | |
| 169 | ||
| 59377 | 170 | text \<open> | 
| 38505 | 171 |   After an interpretation of this locale (say, @{command_def
 | 
| 69505 | 172 |   global_interpretation} \<open>fun_power:\<close> @{term [source] "power (\<lambda>n (f
 | 
| 173 |   :: 'a \<Rightarrow> 'a). f ^^ n)"}), one could naively expect to have a constant \<open>fun_power.powers :: nat list \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a\<close> for which code
 | |
| 37426 | 174 | can be generated. But this not the case: internally, the term | 
| 69505 | 175 | \<open>fun_power.powers\<close> is an abbreviation for the foundational | 
| 37426 | 176 |   term @{term [source] "power.powers (\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)"}
 | 
| 76987 | 177 | (see \<^cite>\<open>"isabelle-locale"\<close> for the details behind). | 
| 37426 | 178 | |
| 61891 
76189756ff65
documentation on last state of the art concerning interpretation
 haftmann parents: 
61779diff
changeset | 179 | Fortunately, a succint solution is available: a dedicated | 
| 61779 | 180 | rewrite definition: | 
| 59377 | 181 | \<close> | 
| 37426 | 182 | |
| 61891 
76189756ff65
documentation on last state of the art concerning interpretation
 haftmann parents: 
61779diff
changeset | 183 | global_interpretation %quote fun_power: power "(\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)" | 
| 61779 | 184 | defines funpows = fun_power.powers | 
| 185 | by unfold_locales | |
| 186 | (simp_all add: fun_eq_iff funpow_mult mult.commute) | |
| 37426 | 187 | |
| 59377 | 188 | text \<open> | 
| 61779 | 189 | \noindent This amends the interpretation morphisms such that | 
| 190 |   occurrences of the foundational term @{term [source] "power.powers (\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)"}
 | |
| 69597 | 191 | are folded to a newly defined constant \<^const>\<open>funpows\<close>. | 
| 37426 | 192 | |
| 193 | After this setup procedure, code generation can continue as usual: | |
| 59377 | 194 | \<close> | 
| 37426 | 195 | |
| 69660 
2bc2a8599369
canonical operation to typeset generated code makes dedicated environment obsolete
 haftmann parents: 
69597diff
changeset | 196 | text %quote \<open> | 
| 69697 
4d95261fab5a
more conventional syntax for code_stmts antiquotation
 haftmann parents: 
69660diff
changeset | 197 |   @{code_stmts funpows constant: Nat.funpow funpows (Haskell)}
 | 
| 61779 | 198 | \<close> | 
| 37426 | 199 | |
| 200 | ||
| 59377 | 201 | subsection \<open>Parallel computation\<close> | 
| 51172 | 202 | |
| 59377 | 203 | text \<open> | 
| 69505 | 204 | Theory \<open>Parallel\<close> in \<^dir>\<open>~~/src/HOL/Library\<close> contains | 
| 51172 | 205 | operations to exploit parallelism inside the Isabelle/ML | 
| 206 | runtime engine. | |
| 59377 | 207 | \<close> | 
| 51172 | 208 | |
| 59377 | 209 | subsection \<open>Imperative data structures\<close> | 
| 28419 | 210 | |
| 59377 | 211 | text \<open> | 
| 38437 | 212 | If you consider imperative data structures as inevitable for a | 
| 213 |   specific application, you should consider \emph{Imperative
 | |
| 214 | Functional Programming with Isabelle/HOL} | |
| 76987 | 215 | \<^cite>\<open>"bulwahn-et-al:2008:imperative"\<close>; the framework described there | 
| 69505 | 216 | is available in session \<open>Imperative_HOL\<close>, together with a | 
| 40752 | 217 | short primer document. | 
| 59377 | 218 | \<close> | 
| 28419 | 219 | |
| 28213 | 220 | end |