src/Tools/code/code_package.ML
author haftmann
Fri, 10 Aug 2007 17:04:34 +0200
changeset 24219 e558fe311376
child 24250 c59c09b09794
permissions -rw-r--r--
new structure for code generator modules
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     1
(*  Title:      Tools/code/code_package.ML
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     2
    ID:         $Id$
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     4
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     5
Code generator translation kernel.  Code generator Isar setup.
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     6
*)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     7
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     8
signature CODE_PACKAGE =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     9
sig
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    10
  (* interfaces *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    11
  val eval_conv: theory
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    12
    -> (CodeThingol.code -> CodeThingol.iterm -> cterm -> thm) -> cterm -> thm;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    13
  val codegen_command: theory -> string -> unit;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    14
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    15
  (* primitive interfaces *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    16
  val eval_term: theory -> (string (*reference name!*) * 'a option ref) * term -> 'a;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    17
  val satisfies_ref: bool option ref;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    18
  val satisfies: theory -> term -> string list -> bool;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    19
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    20
  (* axiomatic interfaces *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    21
  type appgen;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    22
  val add_appconst: string * appgen -> theory -> theory;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    23
  val appgen_let: appgen;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    24
  val appgen_if: appgen;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    25
  val appgen_case: (theory -> term
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    26
    -> ((string * typ) list * ((term * typ) * (term * term) list)) option)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    27
    -> appgen;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    28
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    29
  val timing: bool ref;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    30
end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    31
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    32
structure CodePackage : CODE_PACKAGE =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    33
struct
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    34
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    35
open BasicCodeThingol;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    36
val succeed = CodeThingol.succeed;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    37
val fail = CodeThingol.fail;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    38
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    39
(** code translation **)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    40
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    41
(* theory data *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    42
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    43
type appgen = theory -> ((sort -> sort) * Sorts.algebra) * Consts.T
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    44
  -> CodeFuncgr.T
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    45
  -> (string * typ) * term list -> CodeThingol.transact -> iterm * CodeThingol.transact;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    46
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    47
type appgens = (int * (appgen * stamp)) Symtab.table;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    48
val merge_appgens : appgens * appgens -> appgens =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    49
  Symtab.merge (fn ((bounds1, (_, stamp1)), (bounds2, (_, stamp2))) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    50
    bounds1 = bounds2 andalso stamp1 = stamp2);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    51
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    52
structure Consttab = CodeUnit.Consttab;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    53
type abstypes = typ Symtab.table * CodeUnit.const Consttab.table;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    54
fun merge_abstypes ((typs1, consts1) : abstypes, (typs2, consts2) : abstypes) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    55
  (Symtab.merge (Type.eq_type Vartab.empty) (typs1, typs2),
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    56
    Consttab.merge CodeUnit.eq_const (consts1, consts2));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    57
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    58
structure Translation = TheoryDataFun
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    59
(
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    60
  type T = appgens * abstypes;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    61
  val empty = (Symtab.empty, (Symtab.empty, Consttab.empty));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    62
  val copy = I;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    63
  val extend = I;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    64
  fun merge _ ((appgens1, abstypes1), (appgens2, abstypes2)) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    65
    (merge_appgens (appgens1, appgens2), merge_abstypes (abstypes1, abstypes2));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    66
);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    67
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    68
structure Funcgr = CodeFuncgrRetrieval (fun rewrites thy = []);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    69
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    70
fun code_depgr thy [] = Funcgr.make thy []
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    71
  | code_depgr thy consts =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    72
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    73
        val gr = Funcgr.make thy consts;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    74
        val select = CodeFuncgr.Constgraph.all_succs gr consts;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    75
      in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    76
        CodeFuncgr.Constgraph.subgraph
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    77
          (member CodeUnit.eq_const select) gr
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    78
      end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    79
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    80
fun code_thms thy =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    81
  Pretty.writeln o CodeFuncgr.pretty thy o code_depgr thy;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    82
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    83
fun code_deps thy consts =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    84
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    85
    val gr = code_depgr thy consts;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    86
    fun mk_entry (const, (_, (_, parents))) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    87
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    88
        val name = CodeUnit.string_of_const thy const;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    89
        val nameparents = map (CodeUnit.string_of_const thy) parents;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    90
      in { name = name, ID = name, dir = "", unfold = true,
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    91
        path = "", parents = nameparents }
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    92
      end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    93
    val prgr = CodeFuncgr.Constgraph.fold ((fn x => fn xs => xs @ [x]) o mk_entry) gr [];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    94
  in Present.display_graph prgr end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    95
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    96
structure Program = CodeDataFun
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    97
(
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    98
  type T = CodeThingol.code;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    99
  val empty = CodeThingol.empty_code;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   100
  fun merge _ = CodeThingol.merge_code;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   101
  fun purge _ NONE _ = CodeThingol.empty_code
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   102
    | purge NONE _ _ = CodeThingol.empty_code
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   103
    | purge (SOME thy) (SOME cs) code =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   104
        let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   105
          val cs_exisiting =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   106
            map_filter (CodeName.const_rev thy) (Graph.keys code);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   107
          val dels = (Graph.all_preds code
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   108
              o map (CodeName.const thy)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   109
              o filter (member CodeUnit.eq_const cs_exisiting)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   110
            ) cs;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   111
        in Graph.del_nodes dels code end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   112
);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   113
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   114
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   115
(* translation kernel *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   116
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   117
fun get_abstype thy (tyco, tys) = case Symtab.lookup ((fst o snd o Translation.get) thy) tyco
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   118
 of SOME ty => SOME ((map_atyps (fn TFree (n, _) => nth tys (the (Int.fromString n)))) ty)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   119
  | NONE => NONE;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   120
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   121
fun ensure_def thy = CodeThingol.ensure_def (CodeName.labelled_name thy);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   122
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   123
fun ensure_def_class thy (algbr as ((_, algebra), _)) funcgr class =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   124
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   125
    val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   126
    val (v, cs) = AxClass.params_of_class thy class;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   127
    val class' = CodeName.class thy class;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   128
    val classrels' = map (curry (CodeName.classrel thy) class) superclasses;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   129
    val classops' = map (CodeName.const thy o CodeUnit.const_of_cexpr thy) cs;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   130
    val defgen_class =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   131
      fold_map (ensure_def_class thy algbr funcgr) superclasses
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   132
      ##>> (fold_map (exprgen_typ thy algbr funcgr) o map snd) cs
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   133
      #-> (fn (superclasses, classoptyps) => succeed
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   134
        (CodeThingol.Class (superclasses ~~ classrels', (unprefix "'" v, classops' ~~ classoptyps))))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   135
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   136
    ensure_def thy defgen_class ("generating class " ^ quote class) class'
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   137
    #> pair class'
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   138
  end
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   139
and ensure_def_classrel thy algbr funcgr (subclass, superclass) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   140
  ensure_def_class thy algbr funcgr subclass
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   141
  #>> (fn _ => CodeName.classrel thy (subclass, superclass))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   142
and ensure_def_tyco thy algbr funcgr "fun" trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   143
      trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   144
      |> pair "fun"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   145
  | ensure_def_tyco thy algbr funcgr tyco trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   146
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   147
        fun defgen_datatype trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   148
          let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   149
            val (vs, cos) = Code.get_datatype thy tyco;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   150
          in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   151
            trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   152
            |> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   153
            ||>> fold_map (fn (c, tys) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   154
              fold_map (exprgen_typ thy algbr funcgr) tys
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   155
              #-> (fn tys' =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   156
                pair ((CodeName.const thy o CodeUnit.const_of_cexpr thy)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   157
                  (c, tys ---> Type (tyco, map TFree vs)), tys'))) cos
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   158
            |-> (fn (vs, cos) => succeed (CodeThingol.Datatype (vs, cos)))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   159
          end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   160
        val tyco' = CodeName.tyco thy tyco;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   161
      in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   162
        trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   163
        |> ensure_def thy defgen_datatype ("generating type constructor " ^ quote tyco) tyco'
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   164
        |> pair tyco'
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   165
      end
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   166
and exprgen_tyvar_sort thy (algbr as ((proj_sort, _), _)) funcgr (v, sort) trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   167
  trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   168
  |> fold_map (ensure_def_class thy algbr funcgr) (proj_sort sort)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   169
  |>> (fn sort => (unprefix "'" v, sort))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   170
and exprgen_typ thy algbr funcgr (TFree vs) trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   171
      trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   172
      |> exprgen_tyvar_sort thy algbr funcgr vs
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   173
      |>> (fn (v, sort) => ITyVar v)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   174
  | exprgen_typ thy algbr funcgr (Type (tyco, tys)) trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   175
      case get_abstype thy (tyco, tys)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   176
       of SOME ty =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   177
            trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   178
            |> exprgen_typ thy algbr funcgr ty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   179
        | NONE =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   180
            trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   181
            |> ensure_def_tyco thy algbr funcgr tyco
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   182
            ||>> fold_map (exprgen_typ thy algbr funcgr) tys
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   183
            |>> (fn (tyco, tys) => tyco `%% tys);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   184
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   185
exception CONSTRAIN of (string * typ) * typ;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   186
val timing = ref false;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   187
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   188
fun exprgen_dicts thy (algbr as ((proj_sort, algebra), consts)) funcgr (ty_ctxt, sort_decl) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   189
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   190
    val pp = Sign.pp thy;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   191
    datatype typarg =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   192
        Global of (class * string) * typarg list list
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   193
      | Local of (class * class) list * (string * (int * sort));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   194
    fun class_relation (Global ((_, tyco), yss), _) class =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   195
          Global ((class, tyco), yss)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   196
      | class_relation (Local (classrels, v), subclass) superclass =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   197
          Local ((subclass, superclass) :: classrels, v);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   198
    fun type_constructor tyco yss class =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   199
      Global ((class, tyco), (map o map) fst yss);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   200
    fun type_variable (TFree (v, sort)) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   201
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   202
        val sort' = proj_sort sort;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   203
      in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   204
    val typargs = Sorts.of_sort_derivation pp algebra
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   205
      {class_relation = class_relation, type_constructor = type_constructor,
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   206
       type_variable = type_variable}
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   207
      (ty_ctxt, proj_sort sort_decl);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   208
    fun mk_dict (Global (inst, yss)) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   209
          ensure_def_inst thy algbr funcgr inst
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   210
          ##>> (fold_map o fold_map) mk_dict yss
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   211
          #>> (fn (inst, dss) => DictConst (inst, dss))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   212
      | mk_dict (Local (classrels, (v, (k, sort)))) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   213
          fold_map (ensure_def_classrel thy algbr funcgr) classrels
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   214
          #>> (fn classrels => DictVar (classrels, (unprefix "'" v, (k, length sort))))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   215
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   216
    fold_map mk_dict typargs
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   217
  end
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   218
and exprgen_dict_parms thy (algbr as (_, consts)) funcgr (c, ty_ctxt) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   219
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   220
    val c' = CodeUnit.const_of_cexpr thy (c, ty_ctxt)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   221
    val idf = CodeName.const thy c';
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   222
    val ty_decl = Consts.the_declaration consts idf;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   223
    val (tys, tys_decl) = pairself (curry (Consts.typargs consts) idf) (ty_ctxt, ty_decl);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   224
    val sorts = map (snd o dest_TVar) tys_decl;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   225
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   226
    fold_map (exprgen_dicts thy algbr funcgr) (tys ~~ sorts)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   227
  end
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   228
and ensure_def_inst thy (algbr as ((_, algebra), _)) funcgr (class, tyco) trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   229
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   230
    val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   231
    val (var, classops) = try (AxClass.params_of_class thy) class |> the_default ("'a", [])
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   232
    val vs = Name.names (Name.declare var Name.context) "'a" (Sorts.mg_domain algebra tyco [class]);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   233
    val arity_typ = Type (tyco, map TFree vs);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   234
    fun gen_superarity superclass trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   235
      trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   236
      |> ensure_def_class thy algbr funcgr superclass
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   237
      ||>> ensure_def_classrel thy algbr funcgr (class, superclass)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   238
      ||>> exprgen_dicts thy algbr funcgr (arity_typ, [superclass])
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   239
      |>> (fn ((superclass, classrel), [DictConst (inst, dss)]) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   240
            (superclass, (classrel, (inst, dss))));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   241
    fun gen_classop_def (classop as (c, ty)) trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   242
      trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   243
      |> ensure_def_const thy algbr funcgr (CodeUnit.const_of_cexpr thy classop)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   244
      ||>> exprgen_term thy algbr funcgr (Const (c, map_type_tfree (K arity_typ) ty));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   245
    fun defgen_inst trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   246
      trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   247
      |> ensure_def_class thy algbr funcgr class
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   248
      ||>> ensure_def_tyco thy algbr funcgr tyco
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   249
      ||>> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   250
      ||>> fold_map gen_superarity superclasses
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   251
      ||>> fold_map gen_classop_def classops
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   252
      |-> (fn ((((class, tyco), arity), superarities), classops) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   253
             succeed (CodeThingol.Classinst ((class, (tyco, arity)), (superarities, classops))));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   254
    val inst = CodeName.instance thy (class, tyco);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   255
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   256
    trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   257
    |> ensure_def thy defgen_inst ("generating instance " ^ quote class ^ " / " ^ quote tyco) inst
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   258
    |> pair inst
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   259
  end
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   260
and ensure_def_const thy (algbr as (_, consts)) funcgr (const as (c, opt_tyco)) trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   261
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   262
    val c' = CodeName.const thy const;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   263
    fun defgen_datatypecons trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   264
      trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   265
      |> ensure_def_tyco thy algbr funcgr
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   266
          ((the o Code.get_datatype_of_constr thy) const)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   267
      |-> (fn _ => succeed CodeThingol.Bot);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   268
    fun defgen_classop trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   269
      trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   270
      |> ensure_def_class thy algbr funcgr ((the o AxClass.class_of_param thy) c)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   271
      |-> (fn _ => succeed CodeThingol.Bot);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   272
    fun defgen_fun trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   273
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   274
        val const' = perhaps (Consttab.lookup ((snd o snd o Translation.get) thy)) const;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   275
        val raw_thms = CodeFuncgr.funcs funcgr const';
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   276
        val ty = (Logic.unvarifyT o CodeFuncgr.typ funcgr) const';
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   277
        val thms = if (null o Term.typ_tfrees) ty orelse (null o fst o strip_type) ty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   278
          then raw_thms
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   279
          else map (CodeUnit.expand_eta 1) raw_thms;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   280
        val timeap = if !timing then Output.timeap_msg ("time for " ^ c')
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   281
          else I;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   282
        val msg = cat_lines ("generating code for theorems " :: map string_of_thm thms);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   283
        val vs = (map dest_TFree o Consts.typargs consts) (c', ty);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   284
        val dest_eqthm =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   285
          apfst (snd o strip_comb) o Logic.dest_equals o Logic.unvarify o prop_of;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   286
        fun exprgen_eq (args, rhs) trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   287
          trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   288
          |> fold_map (exprgen_term thy algbr funcgr) args
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   289
          ||>> exprgen_term thy algbr funcgr rhs;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   290
      in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   291
        trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   292
        |> CodeThingol.message msg (fn trns => trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   293
        |> timeap (fold_map (exprgen_eq o dest_eqthm) thms)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   294
        ||>> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   295
        ||>> exprgen_typ thy algbr funcgr ty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   296
        |-> (fn ((eqs, vs), ty) => succeed (CodeThingol.Fun (eqs, (vs, ty)))))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   297
      end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   298
    val defgen = if (is_some o Code.get_datatype_of_constr thy) const
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   299
      then defgen_datatypecons
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   300
      else if is_some opt_tyco
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   301
        orelse (not o is_some o AxClass.class_of_param thy) c
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   302
      then defgen_fun
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   303
      else defgen_classop
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   304
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   305
    trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   306
    |> ensure_def thy defgen ("generating constant " ^ CodeUnit.string_of_const thy const) c'
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   307
    |> pair c'
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   308
  end
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   309
and exprgen_term thy algbr funcgr (Const (c, ty)) trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   310
      trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   311
      |> select_appgen thy algbr funcgr ((c, ty), [])
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   312
  | exprgen_term thy algbr funcgr (Free (v, ty)) trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   313
      trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   314
      |> exprgen_typ thy algbr funcgr ty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   315
      |>> (fn _ => IVar v)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   316
  | exprgen_term thy algbr funcgr (Abs (raw_v, ty, raw_t)) trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   317
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   318
        val (v, t) = Syntax.variant_abs (raw_v, ty, raw_t);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   319
      in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   320
        trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   321
        |> exprgen_typ thy algbr funcgr ty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   322
        ||>> exprgen_term thy algbr funcgr t
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   323
        |>> (fn (ty, t) => (v, ty) `|-> t)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   324
      end
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   325
  | exprgen_term thy algbr funcgr (t as _ $ _) trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   326
      case strip_comb t
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   327
       of (Const (c, ty), ts) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   328
            trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   329
            |> select_appgen thy algbr funcgr ((c, ty), ts)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   330
        | (t', ts) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   331
            trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   332
            |> exprgen_term thy algbr funcgr t'
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   333
            ||>> fold_map (exprgen_term thy algbr funcgr) ts
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   334
            |>> (fn (t, ts) => t `$$ ts)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   335
and appgen_default thy algbr funcgr ((c, ty), ts) trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   336
  trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   337
  |> ensure_def_const thy algbr funcgr (CodeUnit.const_of_cexpr thy (c, ty))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   338
  ||>> fold_map (exprgen_typ thy algbr funcgr) ((fst o Term.strip_type) ty)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   339
  ||>> exprgen_typ thy algbr funcgr ((snd o Term.strip_type) ty)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   340
  ||>> exprgen_dict_parms thy algbr funcgr (c, ty)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   341
  ||>> fold_map (exprgen_term thy algbr funcgr) ts
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   342
  |>> (fn ((((c, tys), ty), iss), ts) => IConst (c, (iss, tys)) `$$ ts)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   343
and select_appgen thy algbr funcgr ((c, ty), ts) trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   344
  case Symtab.lookup (fst (Translation.get thy)) c
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   345
   of SOME (i, (appgen, _)) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   346
        if length ts < i then
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   347
          let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   348
            val k = length ts;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   349
            val tys = (curry Library.take (i - k) o curry Library.drop k o fst o strip_type) ty;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   350
            val ctxt = (fold o fold_aterms)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   351
              (fn Free (v, _) => Name.declare v | _ => I) ts Name.context;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   352
            val vs = Name.names ctxt "a" tys;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   353
          in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   354
            trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   355
            |> fold_map (exprgen_typ thy algbr funcgr) tys
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   356
            ||>> appgen thy algbr funcgr ((c, ty), ts @ map Free vs)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   357
            |>> (fn (tys, t) => map2 (fn (v, _) => pair v) vs tys `|--> t)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   358
          end
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   359
        else if length ts > i then
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   360
          trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   361
          |> appgen thy algbr funcgr ((c, ty), Library.take (i, ts))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   362
          ||>> fold_map (exprgen_term thy algbr funcgr) (Library.drop (i, ts))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   363
          |>> (fn (t, ts) => t `$$ ts)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   364
        else
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   365
          trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   366
          |> appgen thy algbr funcgr ((c, ty), ts)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   367
    | NONE =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   368
        trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   369
        |> appgen_default thy algbr funcgr ((c, ty), ts);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   370
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   371
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   372
(* entrance points into translation kernel *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   373
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   374
fun ensure_def_const' thy algbr funcgr c trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   375
  ensure_def_const thy algbr funcgr c trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   376
  handle CONSTRAIN ((c, ty), ty_decl) => error (
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   377
    "Constant " ^ c ^ " with most general type\n"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   378
    ^ CodeUnit.string_of_typ thy ty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   379
    ^ "\noccurs with type\n"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   380
    ^ CodeUnit.string_of_typ thy ty_decl);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   381
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   382
fun perhaps_def_const thy algbr funcgr c trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   383
  case try (ensure_def_const thy algbr funcgr c) trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   384
   of SOME (c, trns) => (SOME c, trns)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   385
    | NONE => (NONE, trns);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   386
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   387
fun exprgen_term' thy algbr funcgr t trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   388
  exprgen_term thy algbr funcgr t trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   389
  handle CONSTRAIN ((c, ty), ty_decl) => error ("In term " ^ (quote o Sign.string_of_term thy) t
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   390
    ^ ",\nconstant " ^ c ^ " with most general type\n"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   391
    ^ CodeUnit.string_of_typ thy ty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   392
    ^ "\noccurs with type\n"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   393
    ^ CodeUnit.string_of_typ thy ty_decl);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   394
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   395
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   396
(* parametrized application generators, for instantiation in object logic *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   397
(* (axiomatic extensions of translation kernel) *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   398
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   399
fun appgen_case dest_case_expr thy algbr funcgr (app as (c_ty, ts)) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   400
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   401
    val SOME ([], ((st, sty), ds)) = dest_case_expr thy (list_comb (Const c_ty, ts));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   402
    fun clause_gen (dt, bt) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   403
      exprgen_term thy algbr funcgr dt
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   404
      ##>> exprgen_term thy algbr funcgr bt;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   405
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   406
    exprgen_term thy algbr funcgr st
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   407
    ##>> exprgen_typ thy algbr funcgr sty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   408
    ##>> fold_map clause_gen ds
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   409
    ##>> appgen_default thy algbr funcgr app
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   410
    #>> (fn (((se, sty), ds), t0) => ICase (((se, sty), ds), t0))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   411
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   412
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   413
fun appgen_let thy algbr funcgr (app as (_, [st, ct])) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   414
  exprgen_term thy algbr funcgr ct
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   415
  ##>> exprgen_term thy algbr funcgr st
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   416
  ##>> appgen_default thy algbr funcgr app
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   417
  #>> (fn (((v, ty) `|-> be, se), t0) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   418
            ICase (CodeThingol.collapse_let (((v, ty), se), be), t0)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   419
        | (_, t0) => t0);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   420
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   421
fun appgen_if thy algbr funcgr (app as (_, [tb, tt, tf])) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   422
  exprgen_term thy algbr funcgr tb
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   423
  ##>> exprgen_typ thy algbr funcgr (Type ("bool", []))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   424
  ##>> exprgen_term thy algbr funcgr (Const ("True", Type ("bool", [])))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   425
  ##>> exprgen_term thy algbr funcgr tt
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   426
  ##>> exprgen_term thy algbr funcgr (Const ("False", Type ("bool", [])))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   427
  ##>> exprgen_term thy algbr funcgr tf
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   428
  ##>> appgen_default thy algbr funcgr app
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   429
  #>> (fn ((((((tb, B), T), tt), F), tf), t0) => ICase (((tb, B), [(T, tt), (F, tf)]), t0));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   430
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   431
fun add_appconst (c, appgen) thy =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   432
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   433
    val i = (length o fst o strip_type o Sign.the_const_type thy) c;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   434
    val _ = Program.change thy (K CodeThingol.empty_code);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   435
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   436
    (Translation.map o apfst)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   437
      (Symtab.update (c, (i, (appgen, stamp ())))) thy
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   438
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   439
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   440
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   441
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   442
(** abstype and constsubst interface **)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   443
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   444
local
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   445
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   446
fun add_consts thy f (c1, c2 as (c, opt_tyco)) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   447
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   448
    val _ = if
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   449
        is_some (AxClass.class_of_param thy c) andalso is_none opt_tyco
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   450
        orelse is_some (Code.get_datatype_of_constr thy c2)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   451
      then error ("Not a function: " ^ CodeUnit.string_of_const thy c2)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   452
      else ();
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   453
    val funcgr = Funcgr.make thy [c1, c2];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   454
    val ty1 = (f o CodeFuncgr.typ funcgr) c1;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   455
    val ty2 = CodeFuncgr.typ funcgr c2;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   456
    val _ = if Sign.typ_equiv thy (ty1, ty2) then () else
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   457
      error ("Incompatiable type signatures of " ^ CodeUnit.string_of_const thy c1
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   458
        ^ " and " ^ CodeUnit.string_of_const thy c2 ^ ":\n"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   459
        ^ CodeUnit.string_of_typ thy ty1 ^ "\n" ^ CodeUnit.string_of_typ thy ty2);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   460
  in Consttab.update (c1, c2) end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   461
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   462
fun gen_abstyp prep_const prep_typ (raw_abstyp, raw_substtyp) raw_absconsts thy =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   463
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   464
    val abstyp = Type.no_tvars (prep_typ thy raw_abstyp);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   465
    val substtyp = Type.no_tvars (prep_typ thy raw_substtyp);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   466
    val absconsts = (map o pairself) (prep_const thy) raw_absconsts;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   467
    val Type (abstyco, tys) = abstyp handle BIND => error ("Bad type: " ^ Sign.string_of_typ thy abstyp);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   468
    val typarms = map (fst o dest_TFree) tys handle MATCH => error ("Bad type: " ^ Sign.string_of_typ thy abstyp);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   469
    fun mk_index v = 
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   470
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   471
        val k = find_index (fn w => v = w) typarms;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   472
      in if k = ~1
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   473
        then error ("Free type variable: " ^ quote v)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   474
        else TFree (string_of_int k, [])
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   475
      end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   476
    val typpat = map_atyps (fn TFree (v, _) => mk_index v) substtyp;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   477
    fun apply_typpat (Type (tyco, tys)) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   478
          let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   479
            val tys' = map apply_typpat tys;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   480
          in if tyco = abstyco then
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   481
            (map_atyps (fn TFree (n, _) => nth tys' (the (Int.fromString n)))) typpat
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   482
          else
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   483
            Type (tyco, tys')
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   484
          end
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   485
      | apply_typpat ty = ty;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   486
    val _ = Program.change thy (K CodeThingol.empty_code);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   487
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   488
    thy
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   489
    |> (Translation.map o apsnd) (fn (abstypes, abscs) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   490
          (abstypes
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   491
          |> Symtab.update (abstyco, typpat),
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   492
          abscs
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   493
          |> fold (add_consts thy apply_typpat) absconsts)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   494
       )
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   495
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   496
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   497
fun gen_constsubst prep_const raw_constsubsts thy =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   498
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   499
    val constsubsts = (map o pairself) (prep_const thy) raw_constsubsts;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   500
    val _ = Program.change thy (K CodeThingol.empty_code);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   501
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   502
    thy
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   503
    |> (Translation.map o apsnd o apsnd) (fold (add_consts thy I) constsubsts)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   504
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   505
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   506
in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   507
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   508
val abstyp = gen_abstyp (K I) Sign.certify_typ;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   509
val abstyp_e = gen_abstyp CodeUnit.read_const Sign.read_typ;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   510
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   511
val constsubst = gen_constsubst (K I);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   512
val constsubst_e = gen_constsubst CodeUnit.read_const;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   513
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   514
end; (*local*)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   515
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   516
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   517
(** code generation interfaces **)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   518
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   519
(* generic generation combinators *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   520
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   521
fun generate thy funcgr gen it =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   522
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   523
    (*FIXME*)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   524
    val _ = Funcgr.intervene thy funcgr;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   525
    val cs = map_filter (Consttab.lookup ((snd o snd o Translation.get) thy))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   526
      (CodeFuncgr.all funcgr);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   527
    val funcgr' = Funcgr.make thy cs;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   528
    val naming = NameSpace.qualified_names NameSpace.default_naming;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   529
    val consttab = Consts.empty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   530
      |> fold (fn c => Consts.declare naming
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   531
           ((CodeName.const thy c, CodeFuncgr.typ funcgr' c), true))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   532
           (CodeFuncgr.all funcgr');
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   533
    val algbr = (Code.operational_algebra thy, consttab);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   534
  in   
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   535
    Program.change_yield thy
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   536
      (CodeThingol.start_transact (gen thy algbr funcgr' it))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   537
    |> fst
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   538
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   539
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   540
fun eval_conv thy conv =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   541
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   542
    fun conv' funcgr ct =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   543
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   544
        val t = generate thy funcgr exprgen_term' (Thm.term_of ct);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   545
        val consts = CodeThingol.fold_constnames (insert (op =)) t [];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   546
        val code = Program.get thy
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   547
          |> CodeThingol.project_code true [] (SOME consts)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   548
      in conv code t ct end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   549
  in Funcgr.eval_conv thy conv' end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   550
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   551
fun codegen_term thy t =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   552
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   553
    val ct = Thm.cterm_of thy t;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   554
    val (ct', funcgr) = Funcgr.make_term thy (K (K K)) ct;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   555
    val t' = Thm.term_of ct';
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   556
  in generate thy funcgr exprgen_term' t' end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   557
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   558
fun raw_eval_term thy (ref_spec, t) args =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   559
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   560
    val _ = (Term.map_types o Term.map_atyps) (fn _ =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   561
      error ("Term " ^ Sign.string_of_term thy t ^ " contains polymorphic type"))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   562
      t;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   563
    val t' = codegen_term thy t;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   564
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   565
    CodeTarget.eval_term thy CodeName.labelled_name
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   566
      (Program.get thy) (ref_spec, t') args
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   567
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   568
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   569
val satisfies_ref : bool option ref = ref NONE;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   570
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   571
fun eval_term thy t = raw_eval_term thy t [];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   572
fun satisfies thy t witnesses = raw_eval_term thy
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   573
  (("CodePackage.satisfies_ref", satisfies_ref), t) witnesses;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   574
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   575
fun filter_generatable thy consts =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   576
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   577
    val (consts', funcgr) = Funcgr.make_consts thy consts;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   578
    val consts'' = generate thy funcgr (fold_map ooo perhaps_def_const) consts';
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   579
    val consts''' = map_filter (fn (const, SOME _) => SOME const | (_, NONE) => NONE)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   580
      (consts' ~~ consts'');
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   581
  in consts''' end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   582
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   583
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   584
(** toplevel interface and setup **)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   585
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   586
local
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   587
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   588
structure P = OuterParse
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   589
and K = OuterKeyword
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   590
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   591
fun code raw_cs seris thy =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   592
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   593
    val (perm1, cs) = CodeUnit.read_const_exprs thy
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   594
      (filter_generatable thy) raw_cs;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   595
    val (perm2, cs') = case generate thy (Funcgr.make thy cs) (fold_map ooo ensure_def_const') cs
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   596
     of [] => (true, NONE)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   597
      | cs => (false, SOME cs);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   598
    val code = Program.get thy;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   599
    val seris' = map (fn (((target, module), file), args) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   600
      CodeTarget.get_serializer thy target (perm1 orelse perm2) module file args
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   601
        CodeName.labelled_name cs') seris;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   602
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   603
    (map (fn f => f code) seris' : unit list; ())
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   604
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   605
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   606
fun code_thms_cmd thy =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   607
  code_thms thy o snd o CodeUnit.read_const_exprs thy (fst o Funcgr.make_consts thy);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   608
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   609
fun code_deps_cmd thy =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   610
  code_deps thy o snd o CodeUnit.read_const_exprs thy (fst o Funcgr.make_consts thy);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   611
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   612
val (inK, toK, fileK) = ("in", "to", "file");
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   613
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   614
val code_exprP =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   615
  (Scan.repeat P.term
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   616
  -- Scan.repeat (P.$$$ inK |-- P.name
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   617
     -- Scan.option (P.$$$ toK |-- P.name)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   618
     -- Scan.option (P.$$$ fileK |-- P.name)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   619
     -- Scan.optional (P.$$$ "(" |-- P.arguments --| P.$$$ ")") []
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   620
  ) >> (fn (raw_cs, seris) => code raw_cs seris));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   621
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   622
val _ = OuterSyntax.add_keywords [inK, toK, fileK];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   623
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   624
val (codeK, code_abstypeK, code_axiomsK, code_thmsK, code_depsK) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   625
  ("code_gen", "code_abstype", "code_axioms", "code_thms", "code_deps");
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   626
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   627
in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   628
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   629
val codeP =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   630
  OuterSyntax.improper_command codeK "generate executable code for constants"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   631
    K.diag (P.!!! code_exprP >> (fn f => Toplevel.keep (f o Toplevel.theory_of)));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   632
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   633
fun codegen_command thy cmd =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   634
  case Scan.read OuterLex.stopper (P.!!! code_exprP) ((filter OuterLex.is_proper o OuterSyntax.scan) cmd)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   635
   of SOME f => (writeln "Now generating code..."; f thy)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   636
    | NONE => error ("Bad directive " ^ quote cmd);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   637
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   638
val code_abstypeP =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   639
  OuterSyntax.command code_abstypeK "axiomatic abstypes for code generation" K.thy_decl (
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   640
    (P.typ -- P.typ -- Scan.optional (P.$$$ "where" |-- Scan.repeat1
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   641
        (P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.term)) [])
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   642
    >> (Toplevel.theory o uncurry abstyp_e)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   643
  );
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   644
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   645
val code_axiomsP =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   646
  OuterSyntax.command code_axiomsK "axiomatic constant equalities for code generation" K.thy_decl (
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   647
    Scan.repeat1 (P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.term)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   648
    >> (Toplevel.theory o constsubst_e)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   649
  );
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   650
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   651
val code_thmsP =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   652
  OuterSyntax.improper_command code_thmsK "print system of defining equations for code" OuterKeyword.diag
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   653
    (Scan.repeat P.term
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   654
      >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   655
        o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of)));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   656
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   657
val code_depsP =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   658
  OuterSyntax.improper_command code_depsK "visualize dependencies of defining equations for code" OuterKeyword.diag
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   659
    (Scan.repeat P.term
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   660
      >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   661
        o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of)));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   662
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   663
val _ = OuterSyntax.add_parsers [codeP, code_abstypeP, code_axiomsP, code_thmsP, code_depsP];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   664
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   665
end; (* local *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   666
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   667
end; (* struct *)