src/Tools/code/code_package.ML
author haftmann
Tue, 18 Sep 2007 07:36:10 +0200
changeset 24619 c2e6a0f8c30b
parent 24585 c359896d0f48
child 24621 97d403d9ab54
permissions -rw-r--r--
clarified remark
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
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
    10
  (*interfaces*)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    11
  val eval_conv: theory
24381
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
    12
    -> (CodeThingol.code -> CodeThingol.typscheme * CodeThingol.iterm
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
    13
       -> string list -> cterm -> thm)
24283
haftmann
parents: 24250
diff changeset
    14
    -> cterm -> thm;
haftmann
parents: 24250
diff changeset
    15
  val eval_term: theory
24381
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
    16
    -> (CodeThingol.code -> CodeThingol.typscheme * CodeThingol.iterm
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
    17
       -> string list -> cterm -> 'a)
24283
haftmann
parents: 24250
diff changeset
    18
    -> cterm -> 'a;
24585
haftmann
parents: 24436
diff changeset
    19
  val satisfies_ref: (unit -> bool) option ref;
24283
haftmann
parents: 24250
diff changeset
    20
  val satisfies: theory -> cterm -> string list -> bool;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    21
  val codegen_command: theory -> string -> unit;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    22
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
    23
  (*axiomatic interfaces*)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    24
  type appgen;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    25
  val add_appconst: string * appgen -> theory -> theory;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    26
  val appgen_let: appgen;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    27
  val appgen_if: appgen;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    28
  val appgen_case: (theory -> term
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    29
    -> ((string * typ) list * ((term * typ) * (term * term) list)) option)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    30
    -> appgen;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    31
end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    32
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    33
structure CodePackage : CODE_PACKAGE =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    34
struct
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    35
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    36
open BasicCodeThingol;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    37
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    38
(** code translation **)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    39
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    40
(* theory data *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    41
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    42
type appgen = theory -> ((sort -> sort) * Sorts.algebra) * Consts.T
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    43
  -> CodeFuncgr.T
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    44
  -> (string * typ) * term list -> CodeThingol.transact -> iterm * CodeThingol.transact;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    45
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
    46
structure Appgens = TheoryDataFun
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    47
(
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
    48
  type T = (int * (appgen * stamp)) Symtab.table;
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
    49
  val empty = Symtab.empty;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    50
  val copy = I;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    51
  val extend = I;
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
    52
  fun merge _ = Symtab.merge (fn ((bounds1, (_, stamp1)), (bounds2, (_, stamp2))) =>
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
    53
    bounds1 = bounds2 andalso stamp1 = stamp2);
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    54
);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    55
24283
haftmann
parents: 24250
diff changeset
    56
fun code_depgr thy [] = CodeFuncgr.make thy []
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    57
  | code_depgr thy consts =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    58
      let
24283
haftmann
parents: 24250
diff changeset
    59
        val gr = CodeFuncgr.make thy consts;
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
    60
        val select = Graph.all_succs gr consts;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    61
      in
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
    62
        gr
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
    63
        |> Graph.subgraph (member (op =) select) 
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
    64
        |> Graph.map_nodes ((apsnd o map) (Conv.fconv_rule (Class.overload thy)))
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    65
      end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    66
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    67
fun code_thms thy =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    68
  Pretty.writeln o CodeFuncgr.pretty thy o code_depgr 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_deps thy consts =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    71
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    72
    val gr = code_depgr thy consts;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    73
    fun mk_entry (const, (_, (_, parents))) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    74
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    75
        val name = CodeUnit.string_of_const thy const;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    76
        val nameparents = map (CodeUnit.string_of_const thy) parents;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    77
      in { name = name, ID = name, dir = "", unfold = true,
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    78
        path = "", parents = nameparents }
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    79
      end;
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
    80
    val prgr = Graph.fold ((fn x => fn xs => xs @ [x]) o mk_entry) gr [];
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    81
  in Present.display_graph prgr end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    82
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    83
structure Program = CodeDataFun
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    84
(
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    85
  type T = CodeThingol.code;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    86
  val empty = CodeThingol.empty_code;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    87
  fun merge _ = CodeThingol.merge_code;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    88
  fun purge _ NONE _ = CodeThingol.empty_code
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    89
    | purge NONE _ _ = CodeThingol.empty_code
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    90
    | purge (SOME thy) (SOME cs) code =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    91
        let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    92
          val cs_exisiting =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    93
            map_filter (CodeName.const_rev thy) (Graph.keys code);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    94
          val dels = (Graph.all_preds code
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    95
              o map (CodeName.const thy)
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
    96
              o filter (member (op =) cs_exisiting)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    97
            ) cs;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    98
        in Graph.del_nodes dels code end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    99
);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   100
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   101
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   102
(* translation kernel *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   103
24283
haftmann
parents: 24250
diff changeset
   104
val value_name = "Isabelle_Eval.EVAL.EVAL";
haftmann
parents: 24250
diff changeset
   105
haftmann
parents: 24250
diff changeset
   106
fun ensure_def thy = CodeThingol.ensure_def
haftmann
parents: 24250
diff changeset
   107
  (fn s => if s = value_name then "<term>" else CodeName.labelled_name thy s);
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   108
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   109
fun ensure_def_class thy (algbr as ((_, algebra), _)) funcgr class =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   110
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   111
    val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   112
    val (v, cs) = AxClass.params_of_class thy class;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   113
    val class' = CodeName.class thy class;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   114
    val classrels' = map (curry (CodeName.classrel thy) class) superclasses;
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   115
    val classops' = map (CodeName.const thy o fst) cs;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   116
    val defgen_class =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   117
      fold_map (ensure_def_class thy algbr funcgr) superclasses
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   118
      ##>> (fold_map (exprgen_typ thy algbr funcgr) o map snd) cs
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   119
      #>> (fn (superclasses, classoptyps) =>
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   120
        CodeThingol.Class (superclasses ~~ classrels', (unprefix "'" v, classops' ~~ classoptyps)))
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   121
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   122
    ensure_def thy defgen_class ("generating class " ^ quote class) class'
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   123
    #> pair class'
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   124
  end
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   125
and ensure_def_classrel thy algbr funcgr (subclass, superclass) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   126
  ensure_def_class thy algbr funcgr subclass
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   127
  #>> (fn _ => CodeName.classrel thy (subclass, superclass))
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   128
and ensure_def_tyco thy algbr funcgr "fun" =
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   129
      pair "fun"
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   130
  | ensure_def_tyco thy algbr funcgr tyco =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   131
      let
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   132
        val defgen_datatype =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   133
          let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   134
            val (vs, cos) = Code.get_datatype thy tyco;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   135
          in
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   136
            fold_map (exprgen_tyvar_sort thy algbr funcgr) vs
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   137
            ##>> fold_map (fn (c, tys) =>
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   138
              fold_map (exprgen_typ thy algbr funcgr) tys
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   139
              #>> (fn tys' => (CodeName.const thy c, tys'))) cos
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   140
            #>> (fn (vs, cos) => CodeThingol.Datatype (vs, cos))
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   141
          end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   142
        val tyco' = CodeName.tyco thy tyco;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   143
      in
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   144
        ensure_def thy defgen_datatype ("generating type constructor " ^ quote tyco) tyco'
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   145
        #> pair tyco'
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   146
      end
24585
haftmann
parents: 24436
diff changeset
   147
and exprgen_tyvar_sort thy (algbr as ((proj_sort, _), _)) funcgr (v, sort) =
haftmann
parents: 24436
diff changeset
   148
  fold_map (ensure_def_class thy algbr funcgr) (proj_sort sort)
haftmann
parents: 24436
diff changeset
   149
  #>> (fn sort => (unprefix "'" v, sort))
haftmann
parents: 24436
diff changeset
   150
and exprgen_typ thy algbr funcgr (TFree vs) =
haftmann
parents: 24436
diff changeset
   151
      exprgen_tyvar_sort thy algbr funcgr vs
haftmann
parents: 24436
diff changeset
   152
      #>> (fn (v, sort) => ITyVar v)
haftmann
parents: 24436
diff changeset
   153
  | exprgen_typ thy algbr funcgr (Type (tyco, tys)) =
haftmann
parents: 24436
diff changeset
   154
      ensure_def_tyco thy algbr funcgr tyco
haftmann
parents: 24436
diff changeset
   155
      ##>> fold_map (exprgen_typ thy algbr funcgr) tys
haftmann
parents: 24436
diff changeset
   156
      #>> (fn (tyco, tys) => tyco `%% tys);
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   157
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   158
exception CONSTRAIN of (string * typ) * typ;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   159
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   160
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
   161
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   162
    val pp = Sign.pp thy;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   163
    datatype typarg =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   164
        Global of (class * string) * typarg list list
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   165
      | Local of (class * class) list * (string * (int * sort));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   166
    fun class_relation (Global ((_, tyco), yss), _) class =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   167
          Global ((class, tyco), yss)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   168
      | class_relation (Local (classrels, v), subclass) superclass =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   169
          Local ((subclass, superclass) :: classrels, v);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   170
    fun type_constructor tyco yss class =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   171
      Global ((class, tyco), (map o map) fst yss);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   172
    fun type_variable (TFree (v, sort)) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   173
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   174
        val sort' = proj_sort sort;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   175
      in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   176
    val typargs = Sorts.of_sort_derivation pp algebra
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   177
      {class_relation = class_relation, type_constructor = type_constructor,
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   178
       type_variable = type_variable}
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   179
      (ty_ctxt, proj_sort sort_decl);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   180
    fun mk_dict (Global (inst, yss)) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   181
          ensure_def_inst thy algbr funcgr inst
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   182
          ##>> (fold_map o fold_map) mk_dict yss
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   183
          #>> (fn (inst, dss) => DictConst (inst, dss))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   184
      | mk_dict (Local (classrels, (v, (k, sort)))) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   185
          fold_map (ensure_def_classrel thy algbr funcgr) classrels
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   186
          #>> (fn classrels => DictVar (classrels, (unprefix "'" v, (k, length sort))))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   187
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   188
    fold_map mk_dict typargs
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   189
  end
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   190
and exprgen_dict_parms thy (algbr as (_, consts)) funcgr (c, ty_ctxt) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   191
  let
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   192
    val ty_decl = Consts.the_declaration consts c;
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   193
    val (tys, tys_decl) = pairself (curry (Consts.typargs consts) c) (ty_ctxt, ty_decl);
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   194
    val sorts = map (snd o dest_TVar) tys_decl;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   195
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   196
    fold_map (exprgen_dicts thy algbr funcgr) (tys ~~ sorts)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   197
  end
24585
haftmann
parents: 24436
diff changeset
   198
and exprgen_eq thy algbr funcgr thm =
haftmann
parents: 24436
diff changeset
   199
  let
haftmann
parents: 24436
diff changeset
   200
    val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals
haftmann
parents: 24436
diff changeset
   201
      o Logic.unvarify o prop_of) thm;
haftmann
parents: 24436
diff changeset
   202
  in
haftmann
parents: 24436
diff changeset
   203
    fold_map (exprgen_term thy algbr funcgr) args
haftmann
parents: 24436
diff changeset
   204
    ##>> exprgen_term thy algbr funcgr rhs
haftmann
parents: 24436
diff changeset
   205
    #>> rpair thm
haftmann
parents: 24436
diff changeset
   206
  end
haftmann
parents: 24436
diff changeset
   207
and ensure_def_inst thy (algbr as ((_, algebra), _)) funcgr (class, tyco) =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   208
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   209
    val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   210
    val (var, classops) = try (AxClass.params_of_class thy) class |> the_default ("'a", [])
24585
haftmann
parents: 24436
diff changeset
   211
    val vs = Name.names Name.context "'a" (Sorts.mg_domain algebra tyco [class]);
haftmann
parents: 24436
diff changeset
   212
    val sorts' = Sorts.mg_domain (Sign.classes_of thy) tyco [class];
haftmann
parents: 24436
diff changeset
   213
    val vs' = map2 (fn (v, sort1) => fn sort2 => (v,
haftmann
parents: 24436
diff changeset
   214
      Sorts.inter_sort (Sign.classes_of thy) (sort1, sort2))) vs sorts';
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   215
    val arity_typ = Type (tyco, map TFree vs);
24585
haftmann
parents: 24436
diff changeset
   216
    val arity_typ' = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) vs');
haftmann
parents: 24436
diff changeset
   217
    fun gen_superarity superclass =
haftmann
parents: 24436
diff changeset
   218
      ensure_def_class thy algbr funcgr superclass
haftmann
parents: 24436
diff changeset
   219
      ##>> ensure_def_classrel thy algbr funcgr (class, superclass)
haftmann
parents: 24436
diff changeset
   220
      ##>> exprgen_dicts thy algbr funcgr (arity_typ, [superclass])
haftmann
parents: 24436
diff changeset
   221
      #>> (fn ((superclass, classrel), [DictConst (inst, dss)]) =>
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   222
            (superclass, (classrel, (inst, dss))));
24585
haftmann
parents: 24436
diff changeset
   223
    fun gen_classop_def (c, ty) =
haftmann
parents: 24436
diff changeset
   224
      let
haftmann
parents: 24436
diff changeset
   225
        val c_inst = Const (c, map_type_tfree (K arity_typ') ty);
haftmann
parents: 24436
diff changeset
   226
        val thm = Class.unoverload thy (Thm.cterm_of thy c_inst);
haftmann
parents: 24436
diff changeset
   227
        val c_ty = (apsnd Logic.unvarifyT o dest_Const o snd
haftmann
parents: 24436
diff changeset
   228
          o Logic.dest_equals o Thm.prop_of) thm;
haftmann
parents: 24436
diff changeset
   229
      in
haftmann
parents: 24436
diff changeset
   230
        ensure_def_const thy algbr funcgr c
haftmann
parents: 24436
diff changeset
   231
        ##>> exprgen_const thy algbr funcgr c_ty
haftmann
parents: 24436
diff changeset
   232
        #>> (fn (c, IConst c_inst) => ((c, c_inst), thm))
haftmann
parents: 24436
diff changeset
   233
      end;
haftmann
parents: 24436
diff changeset
   234
    val defgen_inst =
haftmann
parents: 24436
diff changeset
   235
      ensure_def_class thy algbr funcgr class
haftmann
parents: 24436
diff changeset
   236
      ##>> ensure_def_tyco thy algbr funcgr tyco
haftmann
parents: 24436
diff changeset
   237
      ##>> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs
haftmann
parents: 24436
diff changeset
   238
      ##>> fold_map gen_superarity superclasses
haftmann
parents: 24436
diff changeset
   239
      ##>> fold_map gen_classop_def classops
haftmann
parents: 24436
diff changeset
   240
      #>> (fn ((((class, tyco), arity), superarities), classops) =>
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   241
             CodeThingol.Classinst ((class, (tyco, arity)), (superarities, classops)));
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   242
    val inst = CodeName.instance thy (class, tyco);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   243
  in
24585
haftmann
parents: 24436
diff changeset
   244
    ensure_def thy defgen_inst ("generating instance " ^ quote class ^ " / " ^ quote tyco) inst
haftmann
parents: 24436
diff changeset
   245
    #> pair inst
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   246
  end
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   247
and ensure_def_const thy (algbr as (_, consts)) funcgr c =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   248
  let
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   249
    val c' = CodeName.const thy c;
24585
haftmann
parents: 24436
diff changeset
   250
    fun defgen_datatypecons tyco =
haftmann
parents: 24436
diff changeset
   251
      ensure_def_tyco thy algbr funcgr tyco
haftmann
parents: 24436
diff changeset
   252
      #>> K CodeThingol.Bot;
haftmann
parents: 24436
diff changeset
   253
    fun defgen_classop class =
haftmann
parents: 24436
diff changeset
   254
      ensure_def_class thy algbr funcgr class
haftmann
parents: 24436
diff changeset
   255
      #>> K CodeThingol.Bot;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   256
    fun defgen_fun trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   257
      let
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   258
        val raw_thms = CodeFuncgr.funcs funcgr c;
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   259
        val ty = (Logic.unvarifyT o CodeFuncgr.typ funcgr) c;
24585
haftmann
parents: 24436
diff changeset
   260
        val vs = (map dest_TFree o Consts.typargs consts) (c, ty);
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   261
        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
   262
          then raw_thms
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   263
          else map (CodeUnit.expand_eta 1) raw_thms;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   264
      in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   265
        trns
24381
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
   266
        |> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   267
        ||>> exprgen_typ thy algbr funcgr ty
24585
haftmann
parents: 24436
diff changeset
   268
        ||>> fold_map (exprgen_eq thy algbr funcgr) thms
24381
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
   269
        |>> (fn ((vs, ty), eqs) => CodeThingol.Fun ((vs, ty), eqs))
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   270
      end;
24585
haftmann
parents: 24436
diff changeset
   271
    val defgen = case Code.get_datatype_of_constr thy c
haftmann
parents: 24436
diff changeset
   272
     of SOME tyco => defgen_datatypecons tyco
haftmann
parents: 24436
diff changeset
   273
      | NONE => (case AxClass.class_of_param thy c
haftmann
parents: 24436
diff changeset
   274
         of SOME class => defgen_classop class
haftmann
parents: 24436
diff changeset
   275
          | NONE => defgen_fun)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   276
  in
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   277
    ensure_def thy defgen ("generating constant " ^ CodeUnit.string_of_const thy c) c'
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   278
    #> pair c'
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   279
  end
24585
haftmann
parents: 24436
diff changeset
   280
and exprgen_term thy algbr funcgr (Const (c, ty)) =
haftmann
parents: 24436
diff changeset
   281
      select_appgen thy algbr funcgr ((c, ty), [])
haftmann
parents: 24436
diff changeset
   282
  | exprgen_term thy algbr funcgr (Free (v, _)) =
haftmann
parents: 24436
diff changeset
   283
      pair (IVar v)
haftmann
parents: 24436
diff changeset
   284
  | exprgen_term thy algbr funcgr (Abs (abs as (_, ty, _))) =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   285
      let
24585
haftmann
parents: 24436
diff changeset
   286
        val (v, t) = Syntax.variant_abs abs;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   287
      in
24585
haftmann
parents: 24436
diff changeset
   288
        exprgen_typ thy algbr funcgr ty
haftmann
parents: 24436
diff changeset
   289
        ##>> exprgen_term thy algbr funcgr t
haftmann
parents: 24436
diff changeset
   290
        #>> (fn (ty, t) => (v, ty) `|-> t)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   291
      end
24585
haftmann
parents: 24436
diff changeset
   292
  | exprgen_term thy algbr funcgr (t as _ $ _) =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   293
      case strip_comb t
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   294
       of (Const (c, ty), ts) =>
24585
haftmann
parents: 24436
diff changeset
   295
            select_appgen thy algbr funcgr ((c, ty), ts)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   296
        | (t', ts) =>
24585
haftmann
parents: 24436
diff changeset
   297
            exprgen_term thy algbr funcgr t'
haftmann
parents: 24436
diff changeset
   298
            ##>> fold_map (exprgen_term thy algbr funcgr) ts
haftmann
parents: 24436
diff changeset
   299
            #>> (fn (t, ts) => t `$$ ts)
haftmann
parents: 24436
diff changeset
   300
and exprgen_const thy algbr funcgr (c, ty) =
haftmann
parents: 24436
diff changeset
   301
  ensure_def_const thy algbr funcgr c
haftmann
parents: 24436
diff changeset
   302
  ##>> exprgen_dict_parms thy algbr funcgr (c, ty)
haftmann
parents: 24436
diff changeset
   303
  ##>> fold_map (exprgen_typ thy algbr funcgr) ((fst o Term.strip_type) ty)
haftmann
parents: 24436
diff changeset
   304
  (*##>> exprgen_typ thy algbr funcgr ((snd o Term.strip_type) ty)*)
haftmann
parents: 24436
diff changeset
   305
  #>> (fn ((c, iss), tys) => IConst (c, (iss, tys)))
haftmann
parents: 24436
diff changeset
   306
and exprgen_app_default thy algbr funcgr (c_ty, ts) =
haftmann
parents: 24436
diff changeset
   307
  exprgen_const thy algbr funcgr c_ty
haftmann
parents: 24436
diff changeset
   308
  ##>> fold_map (exprgen_term thy algbr funcgr) ts
haftmann
parents: 24436
diff changeset
   309
  #>> (fn (t, ts) => t `$$ ts)
haftmann
parents: 24436
diff changeset
   310
and select_appgen thy algbr funcgr ((c, ty), ts) =
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   311
  case Symtab.lookup (Appgens.get thy) c
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   312
   of SOME (i, (appgen, _)) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   313
        if length ts < i then
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   314
          let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   315
            val k = length ts;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   316
            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
   317
            val ctxt = (fold o fold_aterms)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   318
              (fn Free (v, _) => Name.declare v | _ => I) ts Name.context;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   319
            val vs = Name.names ctxt "a" tys;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   320
          in
24585
haftmann
parents: 24436
diff changeset
   321
            fold_map (exprgen_typ thy algbr funcgr) tys
haftmann
parents: 24436
diff changeset
   322
            ##>> appgen thy algbr funcgr ((c, ty), ts @ map Free vs)
haftmann
parents: 24436
diff changeset
   323
            #>> (fn (tys, t) => map2 (fn (v, _) => pair v) vs tys `|--> t)
24219
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
        else if length ts > i then
24585
haftmann
parents: 24436
diff changeset
   326
          appgen thy algbr funcgr ((c, ty), Library.take (i, ts))
haftmann
parents: 24436
diff changeset
   327
          ##>> fold_map (exprgen_term thy algbr funcgr) (Library.drop (i, ts))
haftmann
parents: 24436
diff changeset
   328
          #>> (fn (t, ts) => t `$$ ts)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   329
        else
24585
haftmann
parents: 24436
diff changeset
   330
          appgen thy algbr funcgr ((c, ty), ts)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   331
    | NONE =>
24585
haftmann
parents: 24436
diff changeset
   332
        exprgen_app_default thy algbr funcgr ((c, ty), ts);
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   333
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   334
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   335
(* entrance points into translation kernel *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   336
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   337
fun ensure_def_const' thy algbr funcgr c trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   338
  ensure_def_const thy algbr funcgr c trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   339
  handle CONSTRAIN ((c, ty), ty_decl) => error (
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   340
    "Constant " ^ c ^ " with most general type\n"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   341
    ^ CodeUnit.string_of_typ thy ty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   342
    ^ "\noccurs with type\n"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   343
    ^ CodeUnit.string_of_typ thy ty_decl);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   344
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   345
fun perhaps_def_const thy algbr funcgr c trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   346
  case try (ensure_def_const thy algbr funcgr c) trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   347
   of SOME (c, trns) => (SOME c, trns)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   348
    | NONE => (NONE, trns);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   349
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   350
fun exprgen_term' thy algbr funcgr t trns =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   351
  exprgen_term thy algbr funcgr t trns
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   352
  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
   353
    ^ ",\nconstant " ^ c ^ " with most general type\n"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   354
    ^ CodeUnit.string_of_typ thy ty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   355
    ^ "\noccurs with type\n"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   356
    ^ CodeUnit.string_of_typ thy ty_decl);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   357
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   358
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   359
(* parametrized application generators, for instantiation in object logic *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   360
(* (axiomatic extensions of translation kernel) *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   361
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   362
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
   363
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   364
    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
   365
    fun clause_gen (dt, bt) =
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   366
      exprgen_term thy algbr funcgr
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   367
        (map_aterms (fn Const (c_ty as (c, ty)) => Const
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   368
          (Class.unoverload_const thy c_ty, ty) | t => t) dt)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   369
      ##>> exprgen_term thy algbr funcgr bt;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   370
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   371
    exprgen_term thy algbr funcgr st
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   372
    ##>> exprgen_typ thy algbr funcgr sty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   373
    ##>> fold_map clause_gen ds
24585
haftmann
parents: 24436
diff changeset
   374
    ##>> exprgen_app_default thy algbr funcgr app
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   375
    #>> (fn (((se, sty), ds), t0) => ICase (((se, sty), ds), t0))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   376
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   377
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   378
fun appgen_let thy algbr funcgr (app as (_, [st, ct])) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   379
  exprgen_term thy algbr funcgr ct
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   380
  ##>> exprgen_term thy algbr funcgr st
24585
haftmann
parents: 24436
diff changeset
   381
  ##>> exprgen_app_default thy algbr funcgr app
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   382
  #>> (fn (((v, ty) `|-> be, se), t0) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   383
            ICase (CodeThingol.collapse_let (((v, ty), se), be), t0)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   384
        | (_, t0) => t0);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   385
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   386
fun appgen_if thy algbr funcgr (app as (_, [tb, tt, tf])) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   387
  exprgen_term thy algbr funcgr tb
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   388
  ##>> exprgen_typ thy algbr funcgr (Type ("bool", []))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   389
  ##>> exprgen_term thy algbr funcgr (Const ("True", Type ("bool", [])))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   390
  ##>> exprgen_term thy algbr funcgr tt
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   391
  ##>> exprgen_term thy algbr funcgr (Const ("False", Type ("bool", [])))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   392
  ##>> exprgen_term thy algbr funcgr tf
24585
haftmann
parents: 24436
diff changeset
   393
  ##>> exprgen_app_default thy algbr funcgr app
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   394
  #>> (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
   395
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   396
fun add_appconst (c, appgen) thy =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   397
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   398
    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
   399
    val _ = Program.change thy (K CodeThingol.empty_code);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   400
  in
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   401
    Appgens.map (Symtab.update (c, (i, (appgen, stamp ())))) thy
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   402
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   403
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   404
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   405
(** code generation interfaces **)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   406
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   407
(* generic generation combinators *)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   408
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   409
fun generate thy funcgr gen it =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   410
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   411
    val naming = NameSpace.qualified_names NameSpace.default_naming;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   412
    val consttab = Consts.empty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   413
      |> fold (fn c => Consts.declare naming
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   414
           ((c, CodeFuncgr.typ funcgr c), true))
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   415
           (CodeFuncgr.all funcgr);
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   416
    val algbr = (Code.operational_algebra thy, consttab);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   417
  in   
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   418
    Program.change_yield thy
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24381
diff changeset
   419
      (CodeThingol.start_transact (gen thy algbr funcgr it))
24283
haftmann
parents: 24250
diff changeset
   420
    |> fst
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   421
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   422
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   423
fun code thy permissive cs seris =
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   424
  let
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   425
    val code = Program.get thy;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   426
    val seris' = map (fn (((target, module), file), args) =>
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   427
      CodeTarget.get_serializer thy target permissive module file args
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   428
        CodeName.labelled_name cs) seris;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   429
  in (map (fn f => f code) seris' : unit list; ()) end;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   430
24283
haftmann
parents: 24250
diff changeset
   431
fun raw_eval f thy g =
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   432
  let
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   433
    val value_name = "Isabelle_Eval.EVAL.EVAL";
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   434
    fun ensure_eval thy algbr funcgr t = 
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   435
      let
24381
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
   436
        val ty = fastype_of t;
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   437
        val vs = fold_term_types (K (fold_atyps (insert (eq_fst op =)
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   438
          o dest_TFree))) t [];
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   439
        val defgen_eval =
24381
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
   440
          fold_map (exprgen_tyvar_sort thy algbr funcgr) vs
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
   441
          ##>> exprgen_typ thy algbr funcgr ty
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
   442
          ##>> exprgen_term' thy algbr funcgr t
24585
haftmann
parents: 24436
diff changeset
   443
          #>> (fn ((vs, ty), t) => CodeThingol.Fun ((vs, ty), [(([], t), Drule.dummy_thm)]));
24283
haftmann
parents: 24250
diff changeset
   444
        fun result (dep, code) =
haftmann
parents: 24250
diff changeset
   445
          let
24585
haftmann
parents: 24436
diff changeset
   446
            val CodeThingol.Fun ((vs, ty), [(([], t), _)]) = Graph.get_node code value_name;
24283
haftmann
parents: 24250
diff changeset
   447
            val deps = Graph.imm_succs code value_name;
haftmann
parents: 24250
diff changeset
   448
            val code' = Graph.del_nodes [value_name] code;
haftmann
parents: 24250
diff changeset
   449
            val code'' = CodeThingol.project_code false [] (SOME deps) code';
24381
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
   450
          in ((code'', ((vs, ty), t), deps), (dep, code')) end;
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   451
      in
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   452
        ensure_def thy defgen_eval "evaluation" value_name
24283
haftmann
parents: 24250
diff changeset
   453
        #> result
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   454
      end;
24283
haftmann
parents: 24250
diff changeset
   455
    fun h funcgr ct =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   456
      let
24381
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
   457
        val (code, vs_ty_t, deps) = generate thy funcgr ensure_eval (Thm.term_of ct);
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
   458
      in g code vs_ty_t deps ct end;
24283
haftmann
parents: 24250
diff changeset
   459
  in f thy h end;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   460
24283
haftmann
parents: 24250
diff changeset
   461
fun eval_conv thy = raw_eval CodeFuncgr.eval_conv thy;
haftmann
parents: 24250
diff changeset
   462
fun eval_term thy = raw_eval CodeFuncgr.eval_term thy;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   463
24585
haftmann
parents: 24436
diff changeset
   464
val satisfies_ref : (unit -> bool) option ref = ref NONE;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   465
24283
haftmann
parents: 24250
diff changeset
   466
fun satisfies thy ct witnesses =
haftmann
parents: 24250
diff changeset
   467
  let
24381
560e8ecdf633 improved evaluation interface
haftmann
parents: 24348
diff changeset
   468
    fun evl code ((vs, ty), t) deps ct =
24283
haftmann
parents: 24250
diff changeset
   469
      let
haftmann
parents: 24250
diff changeset
   470
        val t0 = Thm.term_of ct
haftmann
parents: 24250
diff changeset
   471
        val _ = (Term.map_types o Term.map_atyps) (fn _ =>
haftmann
parents: 24250
diff changeset
   472
          error ("Term " ^ Sign.string_of_term thy t0 ^ " contains polymorphic type"))
haftmann
parents: 24250
diff changeset
   473
          t0;
haftmann
parents: 24250
diff changeset
   474
      in
haftmann
parents: 24250
diff changeset
   475
        CodeTarget.eval_term thy ("CodePackage.satisfies_ref", satisfies_ref)
haftmann
parents: 24250
diff changeset
   476
          code (t, ty) witnesses
haftmann
parents: 24250
diff changeset
   477
      end;
haftmann
parents: 24250
diff changeset
   478
  in eval_term thy evl ct end;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   479
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   480
fun filter_generatable thy consts =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   481
  let
24283
haftmann
parents: 24250
diff changeset
   482
    val (consts', funcgr) = CodeFuncgr.make_consts thy consts;
haftmann
parents: 24250
diff changeset
   483
    val consts'' = generate thy funcgr (fold_map ooo perhaps_def_const) consts';
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   484
    val consts''' = map_filter (fn (const, SOME _) => SOME const | (_, NONE) => NONE)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   485
      (consts' ~~ consts'');
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   486
  in consts''' end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   487
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   488
fun generate_const_exprs thy raw_cs =
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   489
  let
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   490
    val (perm1, cs) = CodeUnit.read_const_exprs thy
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   491
      (filter_generatable thy) raw_cs;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   492
    val (perm2, cs') = case generate thy (CodeFuncgr.make thy cs)
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   493
      (fold_map ooo ensure_def_const') cs
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   494
     of [] => (true, NONE)
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   495
      | cs => (false, SOME cs);
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   496
  in (perm1 orelse perm2, cs') end;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   497
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   498
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   499
(** code properties **)
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   500
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   501
fun mk_codeprops thy all_cs sel_cs =
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   502
  let
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   503
    fun select (thmref, thm) = case try (Drule.unvarify o Drule.zero_var_indexes) thm
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   504
     of NONE => NONE
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   505
      | SOME thm => let
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   506
          val t = (ObjectLogic.drop_judgment thy o Thm.prop_of) thm;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   507
          val cs = fold_aterms (fn Const (c, ty) =>
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   508
            cons (Class.unoverload_const thy (c, ty)) | _ => I) t [];
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   509
        in if exists (member (op =) sel_cs) cs
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   510
          andalso forall (member (op =) all_cs) cs
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   511
          then SOME (thmref, thm) else NONE end;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   512
    fun mk_codeprop (thmref, thm) =
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   513
      let
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   514
        val t = ObjectLogic.drop_judgment thy (Thm.prop_of thm);
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   515
        val ty_judg = fastype_of t;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   516
        val tfrees1 = fold_aterms (fn Const (c, ty) =>
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   517
          Term.add_tfreesT ty | _ => I) t [];
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   518
        val vars = Term.add_frees t [];
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   519
        val tfrees2 = fold (Term.add_tfreesT o snd) vars [];
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   520
        val tfrees' = subtract (op =) tfrees2 tfrees1 |> map TFree;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   521
        val ty = map Term.itselfT tfrees' @ map snd vars ---> ty_judg;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   522
        val tfree_vars = map Logic.mk_type tfrees';
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   523
        val c = PureThy.string_of_thmref thmref
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   524
          |> NameSpace.explode
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   525
          |> (fn [x] => [x] | (x::xs) => xs)
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   526
          |> space_implode "_"
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   527
        val propdef = (((c, ty), tfree_vars @ map Free vars), t);
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   528
      in if c = "" then NONE else SOME (thmref, propdef) end;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   529
  in
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   530
    PureThy.thms_containing thy ([], [])
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   531
    |> maps PureThy.selections
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   532
    |> map_filter select
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   533
    |> map_filter mk_codeprop
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   534
  end;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   535
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   536
fun add_codeprops all_cs sel_cs thy =
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   537
  let
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   538
    val codeprops = mk_codeprops thy all_cs sel_cs;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   539
    fun lift_name_yield f x = (Name.context, x) |> f ||> snd;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   540
    fun add (thmref, (((raw_c, ty), ts), t)) (names, thy) =
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   541
      let
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   542
        val _ = warning ("Adding theorem " ^ quote (PureThy.string_of_thmref thmref)
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   543
          ^ " as code property " ^ quote raw_c);
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   544
        val ([raw_c'], names') = Name.variants [raw_c] names;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   545
      in
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   546
        thy
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   547
        |> PureThy.simple_def ("", []) (((raw_c', ty, Syntax.NoSyn), ts), t)
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   548
        ||> pair names'
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   549
      end;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   550
  in
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   551
    thy
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   552
    |> Sign.sticky_prefix "codeprop"
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   553
    |> lift_name_yield (fold_map add codeprops)
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   554
    ||> Sign.restore_naming thy
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   555
    |-> (fn c_thms => fold (Code.add_func true o snd) c_thms #> pair c_thms)
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   556
  end;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   557
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   558
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   559
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   560
(** toplevel interface and setup **)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   561
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   562
local
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   563
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   564
structure P = OuterParse
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   565
and K = OuterKeyword
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   566
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   567
fun code_cmd raw_cs seris thy =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   568
  let
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   569
    val (permissive, cs) = generate_const_exprs thy raw_cs;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   570
    val _ = code thy permissive cs seris;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   571
  in () end;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   572
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   573
fun code_thms_cmd thy =
24283
haftmann
parents: 24250
diff changeset
   574
  code_thms thy o snd o CodeUnit.read_const_exprs thy (fst o CodeFuncgr.make_consts thy);
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   575
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   576
fun code_deps_cmd thy =
24283
haftmann
parents: 24250
diff changeset
   577
  code_deps thy o snd o CodeUnit.read_const_exprs thy (fst o CodeFuncgr.make_consts thy);
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   578
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   579
fun code_props_cmd raw_cs seris thy =
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   580
  let
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   581
    val (_, all_cs) = generate_const_exprs thy ["*"];
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   582
    val (permissive, cs) = generate_const_exprs thy raw_cs;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   583
    val (c_thms, thy') = add_codeprops (map (the o CodeName.const_rev thy) (these all_cs))
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   584
      (map (the o CodeName.const_rev thy) (these cs)) thy;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   585
    val prop_cs = (filter_generatable thy' o map fst) c_thms;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   586
    val _ = if null seris then [] else generate thy' (CodeFuncgr.make thy' prop_cs)
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   587
      (fold_map ooo ensure_def_const') prop_cs;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   588
    val _ = if null seris then () else code thy' permissive
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   589
      (SOME (map (CodeName.const thy') prop_cs)) seris;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   590
  in thy' end;
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   591
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   592
val (inK, module_nameK, fileK) = ("in", "module_name", "file");
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   593
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   594
fun code_exprP cmd =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   595
  (Scan.repeat P.term
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   596
  -- Scan.repeat (P.$$$ inK |-- P.name
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   597
     -- Scan.option (P.$$$ module_nameK |-- P.name)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   598
     -- Scan.option (P.$$$ fileK |-- P.name)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   599
     -- Scan.optional (P.$$$ "(" |-- P.arguments --| P.$$$ ")") []
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   600
  ) >> (fn (raw_cs, seris) => cmd raw_cs seris));
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   601
24250
c59c09b09794 *** empty log message ***
haftmann
parents: 24219
diff changeset
   602
val _ = OuterSyntax.add_keywords [inK, module_nameK, fileK];
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   603
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   604
val (codeK, code_thmsK, code_depsK, code_propsK) =
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   605
  ("export_code", "code_thms", "code_deps", "code_props");
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   606
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   607
in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   608
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   609
val codeP =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   610
  OuterSyntax.improper_command codeK "generate executable code for constants"
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   611
    K.diag (P.!!! (code_exprP code_cmd) >> (fn f => Toplevel.keep (f o Toplevel.theory_of)));
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   612
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   613
fun codegen_command thy cmd =
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   614
  case Scan.read OuterLex.stopper (P.!!! (code_exprP code_cmd)) ((filter OuterLex.is_proper o OuterSyntax.scan) cmd)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   615
   of SOME f => (writeln "Now generating code..."; f thy)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   616
    | NONE => error ("Bad directive " ^ quote cmd);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   617
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   618
val code_thmsP =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   619
  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
   620
    (Scan.repeat P.term
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   621
      >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   622
        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
   623
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   624
val code_depsP =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   625
  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
   626
    (Scan.repeat P.term
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   627
      >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   628
        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
   629
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   630
val code_propsP =
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   631
  OuterSyntax.command code_propsK "generate characteristic properties for executable constants"
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   632
    K.thy_decl (P.!!! (code_exprP code_props_cmd) >> Toplevel.theory);
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   633
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   634
val _ = OuterSyntax.add_parsers [codeP, code_thmsP, code_depsP, code_propsP];
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   635
24436
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   636
end; (*local*)
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   637
b694324cd2be added code_props
haftmann
parents: 24423
diff changeset
   638
end; (*struct*)