| 28213 |      1 | theory "ML"
 | 
|  |      2 | imports Setup
 | 
|  |      3 | begin
 | 
|  |      4 | 
 | 
| 28419 |      5 | section {* ML system interfaces \label{sec:ml} *}
 | 
|  |      6 | 
 | 
|  |      7 | text {*
 | 
|  |      8 |   Since the code generator framework not only aims to provide
 | 
|  |      9 |   a nice Isar interface but also to form a base for
 | 
|  |     10 |   code-generation-based applications, here a short
 | 
|  |     11 |   description of the most important ML interfaces.
 | 
|  |     12 | *}
 | 
|  |     13 | 
 | 
|  |     14 | subsection {* Executable theory content: @{text Code} *}
 | 
|  |     15 | 
 | 
|  |     16 | text {*
 | 
|  |     17 |   This Pure module implements the core notions of
 | 
|  |     18 |   executable content of a theory.
 | 
|  |     19 | *}
 | 
|  |     20 | 
 | 
|  |     21 | subsubsection {* Managing executable content *}
 | 
|  |     22 | 
 | 
|  |     23 | text %mlref {*
 | 
|  |     24 |   \begin{mldecls}
 | 
|  |     25 |   @{index_ML Code.add_eqn: "thm -> theory -> theory"} \\
 | 
|  |     26 |   @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\
 | 
| 31143 |     27 |   @{index_ML Code_Preproc.map_pre: "(simpset -> simpset) -> theory -> theory"} \\
 | 
|  |     28 |   @{index_ML Code_Preproc.map_post: "(simpset -> simpset) -> theory -> theory"} \\
 | 
|  |     29 |   @{index_ML Code_Preproc.add_functrans: "string * (theory -> (thm * bool) list -> (thm * bool) list option)
 | 
| 28419 |     30 |     -> theory -> theory"} \\
 | 
| 31143 |     31 |   @{index_ML Code_Preproc.del_functrans: "string -> theory -> theory"} \\
 | 
| 28419 |     32 |   @{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\
 | 
|  |     33 |   @{index_ML Code.get_datatype: "theory -> string
 | 
|  |     34 |     -> (string * sort) list * (string * typ list) list"} \\
 | 
|  |     35 |   @{index_ML Code.get_datatype_of_constr: "theory -> string -> string option"}
 | 
|  |     36 |   \end{mldecls}
 | 
|  |     37 | 
 | 
|  |     38 |   \begin{description}
 | 
|  |     39 | 
 | 
|  |     40 |   \item @{ML Code.add_eqn}~@{text "thm"}~@{text "thy"} adds function
 | 
|  |     41 |      theorem @{text "thm"} to executable content.
 | 
|  |     42 | 
 | 
|  |     43 |   \item @{ML Code.del_eqn}~@{text "thm"}~@{text "thy"} removes function
 | 
|  |     44 |      theorem @{text "thm"} from executable content, if present.
 | 
|  |     45 | 
 | 
| 31143 |     46 |   \item @{ML Code_Preproc.map_pre}~@{text "f"}~@{text "thy"} changes
 | 
| 28419 |     47 |      the preprocessor simpset.
 | 
|  |     48 | 
 | 
| 31143 |     49 |   \item @{ML Code_Preproc.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds
 | 
| 28419 |     50 |      function transformer @{text f} (named @{text name}) to executable content;
 | 
| 29560 |     51 |      @{text f} is a transformer of the code equations belonging
 | 
| 28419 |     52 |      to a certain function definition, depending on the
 | 
|  |     53 |      current theory context.  Returning @{text NONE} indicates that no
 | 
|  |     54 |      transformation took place;  otherwise, the whole process will be iterated
 | 
| 29560 |     55 |      with the new code equations.
 | 
| 28419 |     56 | 
 | 
| 31143 |     57 |   \item @{ML Code_Preproc.del_functrans}~@{text "name"}~@{text "thy"} removes
 | 
| 28419 |     58 |      function transformer named @{text name} from executable content.
 | 
|  |     59 | 
 | 
|  |     60 |   \item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds
 | 
|  |     61 |      a datatype to executable content, with generation
 | 
|  |     62 |      set @{text cs}.
 | 
|  |     63 | 
 | 
|  |     64 |   \item @{ML Code.get_datatype_of_constr}~@{text "thy"}~@{text "const"}
 | 
|  |     65 |      returns type constructor corresponding to
 | 
|  |     66 |      constructor @{text const}; returns @{text NONE}
 | 
|  |     67 |      if @{text const} is no constructor.
 | 
|  |     68 | 
 | 
|  |     69 |   \end{description}
 | 
|  |     70 | *}
 | 
|  |     71 | 
 | 
|  |     72 | subsection {* Auxiliary *}
 | 
|  |     73 | 
 | 
|  |     74 | text %mlref {*
 | 
|  |     75 |   \begin{mldecls}
 | 
| 31999 |     76 |   @{index_ML Code.read_const: "theory -> string -> string"}
 | 
| 28419 |     77 |   \end{mldecls}
 | 
|  |     78 | 
 | 
|  |     79 |   \begin{description}
 | 
|  |     80 | 
 | 
| 31156 |     81 |   \item @{ML Code.read_const}~@{text thy}~@{text s}
 | 
| 28419 |     82 |      reads a constant as a concrete term expression @{text s}.
 | 
|  |     83 | 
 | 
|  |     84 |   \end{description}
 | 
|  |     85 | 
 | 
|  |     86 | *}
 | 
|  |     87 | 
 | 
|  |     88 | subsection {* Implementing code generator applications *}
 | 
|  |     89 | 
 | 
|  |     90 | text {*
 | 
|  |     91 |   Implementing code generator applications on top
 | 
|  |     92 |   of the framework set out so far usually not only
 | 
|  |     93 |   involves using those primitive interfaces
 | 
|  |     94 |   but also storing code-dependent data and various
 | 
|  |     95 |   other things.
 | 
|  |     96 | *}
 | 
|  |     97 | 
 | 
|  |     98 | subsubsection {* Data depending on the theory's executable content *}
 | 
|  |     99 | 
 | 
|  |    100 | text {*
 | 
|  |    101 |   Due to incrementality of code generation, changes in the
 | 
|  |    102 |   theory's executable content have to be propagated in a
 | 
|  |    103 |   certain fashion.  Additionally, such changes may occur
 | 
|  |    104 |   not only during theory extension but also during theory
 | 
|  |    105 |   merge, which is a little bit nasty from an implementation
 | 
|  |    106 |   point of view.  The framework provides a solution
 | 
|  |    107 |   to this technical challenge by providing a functorial
 | 
|  |    108 |   data slot @{ML_functor CodeDataFun}; on instantiation
 | 
|  |    109 |   of this functor, the following types and operations
 | 
|  |    110 |   are required:
 | 
|  |    111 | 
 | 
|  |    112 |   \medskip
 | 
|  |    113 |   \begin{tabular}{l}
 | 
|  |    114 |   @{text "type T"} \\
 | 
|  |    115 |   @{text "val empty: T"} \\
 | 
| 28635 |    116 |   @{text "val purge: theory \<rightarrow> string list option \<rightarrow> T \<rightarrow> T"}
 | 
| 28419 |    117 |   \end{tabular}
 | 
|  |    118 | 
 | 
|  |    119 |   \begin{description}
 | 
|  |    120 | 
 | 
|  |    121 |   \item @{text T} the type of data to store.
 | 
|  |    122 | 
 | 
|  |    123 |   \item @{text empty} initial (empty) data.
 | 
|  |    124 | 
 | 
|  |    125 |   \item @{text purge}~@{text thy}~@{text consts} propagates changes in executable content;
 | 
|  |    126 |     @{text consts} indicates the kind
 | 
|  |    127 |     of change: @{ML NONE} stands for a fundamental change
 | 
|  |    128 |     which invalidates any existing code, @{text "SOME consts"}
 | 
|  |    129 |     hints that executable content for constants @{text consts}
 | 
|  |    130 |     has changed.
 | 
|  |    131 | 
 | 
|  |    132 |   \end{description}
 | 
|  |    133 | 
 | 
| 28447 |    134 |   \noindent An instance of @{ML_functor CodeDataFun} provides the following
 | 
| 28419 |    135 |   interface:
 | 
|  |    136 | 
 | 
|  |    137 |   \medskip
 | 
|  |    138 |   \begin{tabular}{l}
 | 
|  |    139 |   @{text "get: theory \<rightarrow> T"} \\
 | 
|  |    140 |   @{text "change: theory \<rightarrow> (T \<rightarrow> T) \<rightarrow> T"} \\
 | 
|  |    141 |   @{text "change_yield: theory \<rightarrow> (T \<rightarrow> 'a * T) \<rightarrow> 'a * T"}
 | 
|  |    142 |   \end{tabular}
 | 
|  |    143 | 
 | 
|  |    144 |   \begin{description}
 | 
|  |    145 | 
 | 
|  |    146 |   \item @{text get} retrieval of the current data.
 | 
|  |    147 | 
 | 
|  |    148 |   \item @{text change} update of current data (cached!)
 | 
|  |    149 |     by giving a continuation.
 | 
|  |    150 | 
 | 
|  |    151 |   \item @{text change_yield} update with side result.
 | 
|  |    152 | 
 | 
|  |    153 |   \end{description}
 | 
|  |    154 | *}
 | 
|  |    155 | 
 | 
|  |    156 | text {*
 | 
| 28447 |    157 |   \bigskip
 | 
|  |    158 | 
 | 
| 28419 |    159 |   \emph{Happy proving, happy hacking!}
 | 
|  |    160 | *}
 | 
| 28213 |    161 | 
 | 
|  |    162 | end
 |