src/HOL/Tools/datatype_codegen.ML
author haftmann
Fri, 21 Jul 2006 14:47:22 +0200
changeset 20177 0af885e3dabf
parent 20105 454f4be984b7
child 20182 79c9ff40d760
permissions -rw-r--r--
hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12446
86887b40aeb1 Tuned header.
berghofe
parents: 12445
diff changeset
     1
(*  Title:      HOL/datatype_codegen.ML
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
     2
    ID:         $Id$
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
     3
    Author:     Stefan Berghofer & Florian Haftmann, TU Muenchen
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
     4
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
     5
Code generator for inductive datatypes.
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
     6
*)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
     7
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
     8
signature DATATYPE_CODEGEN =
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
     9
sig
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
    10
  val get_datatype_spec_thms: theory -> string
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
    11
    -> (((string * sort) list * (string * typ list) list) * tactic) option
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
    12
  val get_all_datatype_cons: theory -> (string * string) list
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
    13
  val dest_case_expr: theory -> term
20177
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
    14
    -> ((string * typ) list * ((term * typ) * (term * term) list)) option
19886
6bec6daac280 added hook for case const defs
haftmann
parents: 19818
diff changeset
    15
  val add_datatype_case_const: string -> theory -> theory
6bec6daac280 added hook for case const defs
haftmann
parents: 19818
diff changeset
    16
  val add_datatype_case_defs: string -> theory -> theory
20177
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
    17
  val datatypes_dependency: theory -> string list list
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
    18
  val get_datatype_mut_specs: theory -> string list
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
    19
    -> ((string * sort) list * (string * (string * typ list) list) list)
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
    20
  val get_datatype_arities: theory -> string list -> sort
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
    21
    -> (string * (((string * sort list) * sort)  * term list)) list option
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
    22
  val datatype_prove_arities : tactic -> string list -> sort
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
    23
    -> ((string * term list) list
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
    24
    -> ((bstring * attribute list) * term) list) -> theory -> theory
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18702
diff changeset
    25
  val setup: theory -> theory
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    26
end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    27
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    28
structure DatatypeCodegen : DATATYPE_CODEGEN =
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    29
struct
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    30
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    31
open Codegen;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    32
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    33
fun mk_tuple [p] = p
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    34
  | mk_tuple ps = Pretty.block (Pretty.str "(" ::
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    35
      List.concat (separate [Pretty.str ",", Pretty.brk 1] (map single ps)) @
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    36
        [Pretty.str ")"]);
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    37
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    38
(**** datatype definition ****)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    39
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    40
(* find shortest path to constructor with no recursive arguments *)
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    41
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    42
fun find_nonempty (descr: DatatypeAux.descr) is i =
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    43
  let
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
    44
    val (_, _, constrs) = valOf (AList.lookup (op =) descr i);
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    45
    fun arg_nonempty (_, DatatypeAux.DtRec i) = if i mem is then NONE
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    46
          else Option.map (curry op + 1 o snd) (find_nonempty descr (i::is) i)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    47
      | arg_nonempty _ = SOME 0;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    48
    fun max xs = Library.foldl
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    49
      (fn (NONE, _) => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    50
        | (SOME i, SOME j) => SOME (Int.max (i, j))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    51
        | (_, NONE) => NONE) (SOME 0, xs);
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    52
    val xs = sort (int_ord o pairself snd)
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    53
      (List.mapPartial (fn (s, dts) => Option.map (pair s)
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    54
        (max (map (arg_nonempty o DatatypeAux.strip_dtyp) dts))) constrs)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    55
  in case xs of [] => NONE | x :: _ => SOME x end;
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    56
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    57
fun add_dt_defs thy defs dep module gr (descr: DatatypeAux.descr) =
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    58
  let
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    59
    val sg = sign_of thy;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    60
    val tab = DatatypePackage.get_datatypes thy;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    61
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    62
    val descr' = List.filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr;
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    63
    val rtnames = map (#1 o snd) (List.filter (fn (_, (_, _, cs)) =>
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    64
      exists (exists DatatypeAux.is_rec_type o snd) cs) descr');
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    65
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    66
    val (_, (tname, _, _)) :: _ = descr';
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    67
    val node_id = tname ^ " (type)";
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    68
    val module' = if_library (thyname_of_type tname thy) module;
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    69
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    70
    fun mk_dtdef gr prfx [] = (gr, [])
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    71
      | mk_dtdef gr prfx ((_, (tname, dts, cs))::xs) =
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    72
          let
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    73
            val tvs = map DatatypeAux.dest_DtTFree dts;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    74
            val sorts = map (rpair []) tvs;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    75
            val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs;
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    76
            val (gr', (_, type_id)) = mk_type_id module' tname gr;
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    77
            val (gr'', ps) =
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    78
              foldl_map (fn (gr, (cname, cargs)) =>
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    79
                foldl_map (invoke_tycodegen thy defs node_id module' false)
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    80
                  (gr, cargs) |>>>
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    81
                mk_const_id module' cname) (gr', cs');
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    82
            val (gr''', rest) = mk_dtdef gr'' "and " xs
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    83
          in
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    84
            (gr''',
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    85
             Pretty.block (Pretty.str prfx ::
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    86
               (if null tvs then [] else
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    87
                  [mk_tuple (map Pretty.str tvs), Pretty.str " "]) @
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    88
               [Pretty.str (type_id ^ " ="), Pretty.brk 1] @
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    89
               List.concat (separate [Pretty.brk 1, Pretty.str "| "]
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    90
                 (map (fn (ps', (_, cname)) => [Pretty.block
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    91
                   (Pretty.str cname ::
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    92
                    (if null ps' then [] else
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    93
                     List.concat ([Pretty.str " of", Pretty.brk 1] ::
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    94
                       separate [Pretty.str " *", Pretty.brk 1]
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    95
                         (map single ps'))))]) ps))) :: rest)
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
    96
          end;
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
    97
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    98
    fun mk_term_of_def gr prfx [] = []
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    99
      | mk_term_of_def gr prfx ((_, (tname, dts, cs)) :: xs) =
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   100
          let
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   101
            val tvs = map DatatypeAux.dest_DtTFree dts;
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   102
            val sorts = map (rpair []) tvs;
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   103
            val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs;
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   104
            val dts' = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   105
            val T = Type (tname, dts');
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   106
            val rest = mk_term_of_def gr "and " xs;
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   107
            val (_, eqs) = foldl_map (fn (prfx, (cname, Ts)) =>
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   108
              let val args = map (fn i =>
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   109
                Pretty.str ("x" ^ string_of_int i)) (1 upto length Ts)
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   110
              in ("  | ", Pretty.blk (4,
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   111
                [Pretty.str prfx, mk_term_of gr module' false T, Pretty.brk 1,
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   112
                 if null Ts then Pretty.str (snd (get_const_id cname gr))
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   113
                 else parens (Pretty.block
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   114
                   [Pretty.str (snd (get_const_id cname gr)),
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   115
                    Pretty.brk 1, mk_tuple args]),
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   116
                 Pretty.str " =", Pretty.brk 1] @
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   117
                 List.concat (separate [Pretty.str " $", Pretty.brk 1]
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   118
                   ([Pretty.str ("Const (\"" ^ cname ^ "\","), Pretty.brk 1,
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   119
                     mk_type false (Ts ---> T), Pretty.str ")"] ::
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   120
                    map (fn (x, U) => [Pretty.block [mk_term_of gr module' false U,
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   121
                      Pretty.brk 1, x]]) (args ~~ Ts)))))
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   122
              end) (prfx, cs')
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   123
          in eqs @ rest end;
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   124
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   125
    fun mk_gen_of_def gr prfx [] = []
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   126
      | mk_gen_of_def gr prfx ((i, (tname, dts, cs)) :: xs) =
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   127
          let
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   128
            val tvs = map DatatypeAux.dest_DtTFree dts;
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   129
            val sorts = map (rpair []) tvs;
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   130
            val (cs1, cs2) =
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   131
              List.partition (exists DatatypeAux.is_rec_type o snd) cs;
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
   132
            val SOME (cname, _) = find_nonempty descr [i] i;
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   133
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   134
            fun mk_delay p = Pretty.block
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   135
              [Pretty.str "fn () =>", Pretty.brk 1, p];
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   136
15397
5260ac75e07c Fixed bug in mk_gen_of_def that could cause non-termination of the generator
berghofe
parents: 14981
diff changeset
   137
            fun mk_constr s b (cname, dts) =
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   138
              let
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   139
                val gs = map (fn dt => mk_app false (mk_gen gr module' false rtnames s
15397
5260ac75e07c Fixed bug in mk_gen_of_def that could cause non-termination of the generator
berghofe
parents: 14981
diff changeset
   140
                    (DatatypeAux.typ_of_dtyp descr sorts dt))
5260ac75e07c Fixed bug in mk_gen_of_def that could cause non-termination of the generator
berghofe
parents: 14981
diff changeset
   141
                  [Pretty.str (if b andalso DatatypeAux.is_rec_type dt then "0"
5260ac75e07c Fixed bug in mk_gen_of_def that could cause non-termination of the generator
berghofe
parents: 14981
diff changeset
   142
                     else "j")]) dts;
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   143
                val (_, id) = get_const_id cname gr
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   144
              in case gs of
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   145
                  _ :: _ :: _ => Pretty.block
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   146
                    [Pretty.str id, Pretty.brk 1, mk_tuple gs]
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   147
                | _ => mk_app false (Pretty.str id) (map parens gs)
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   148
              end;
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   149
15397
5260ac75e07c Fixed bug in mk_gen_of_def that could cause non-termination of the generator
berghofe
parents: 14981
diff changeset
   150
            fun mk_choice [c] = mk_constr "(i-1)" false c
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   151
              | mk_choice cs = Pretty.block [Pretty.str "one_of",
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   152
                  Pretty.brk 1, Pretty.blk (1, Pretty.str "[" ::
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   153
                  List.concat (separate [Pretty.str ",", Pretty.fbrk]
15397
5260ac75e07c Fixed bug in mk_gen_of_def that could cause non-termination of the generator
berghofe
parents: 14981
diff changeset
   154
                    (map (single o mk_delay o mk_constr "(i-1)" false) cs)) @
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   155
                  [Pretty.str "]"]), Pretty.brk 1, Pretty.str "()"];
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   156
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   157
            val gs = map (Pretty.str o suffix "G" o strip_tname) tvs;
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   158
            val gen_name = "gen_" ^ snd (get_type_id tname gr)
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   159
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   160
          in
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   161
            Pretty.blk (4, separate (Pretty.brk 1) 
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   162
                (Pretty.str (prfx ^ gen_name ^
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   163
                   (if null cs1 then "" else "'")) :: gs @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   164
                 (if null cs1 then [] else [Pretty.str "i"]) @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   165
                 [Pretty.str "j"]) @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   166
              [Pretty.str " =", Pretty.brk 1] @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   167
              (if not (null cs1) andalso not (null cs2)
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   168
               then [Pretty.str "frequency", Pretty.brk 1,
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   169
                 Pretty.blk (1, [Pretty.str "[",
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   170
                   mk_tuple [Pretty.str "i", mk_delay (mk_choice cs1)],
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   171
                   Pretty.str ",", Pretty.fbrk,
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   172
                   mk_tuple [Pretty.str "1", mk_delay (mk_choice cs2)],
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   173
                   Pretty.str "]"]), Pretty.brk 1, Pretty.str "()"]
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   174
               else if null cs2 then
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   175
                 [Pretty.block [Pretty.str "(case", Pretty.brk 1,
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   176
                   Pretty.str "i", Pretty.brk 1, Pretty.str "of",
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   177
                   Pretty.brk 1, Pretty.str "0 =>", Pretty.brk 1,
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   178
                   mk_constr "0" true (cname, valOf (AList.lookup (op =) cs cname)),
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   179
                   Pretty.brk 1, Pretty.str "| _ =>", Pretty.brk 1,
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   180
                   mk_choice cs1, Pretty.str ")"]]
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   181
               else [mk_choice cs2])) ::
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   182
            (if null cs1 then []
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   183
             else [Pretty.blk (4, separate (Pretty.brk 1) 
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   184
                 (Pretty.str ("and " ^ gen_name) :: gs @ [Pretty.str "i"]) @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   185
               [Pretty.str " =", Pretty.brk 1] @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   186
               separate (Pretty.brk 1) (Pretty.str (gen_name ^ "'") :: gs @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   187
                 [Pretty.str "i", Pretty.str "i"]))]) @
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   188
            mk_gen_of_def gr "and " xs
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   189
          end
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   190
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   191
  in
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   192
    (add_edge_acyclic (node_id, dep) gr
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   193
        handle Graph.CYCLES _ => gr) handle Graph.UNDEF _ =>
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   194
         let
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   195
           val gr1 = add_edge (node_id, dep)
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   196
             (new_node (node_id, (NONE, "", "")) gr);
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   197
           val (gr2, dtdef) = mk_dtdef gr1 "datatype " descr';
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   198
         in
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   199
           map_node node_id (K (NONE, module',
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   200
             Pretty.string_of (Pretty.blk (0, separate Pretty.fbrk dtdef @
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   201
               [Pretty.str ";"])) ^ "\n\n" ^
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   202
             (if "term_of" mem !mode then
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   203
                Pretty.string_of (Pretty.blk (0, separate Pretty.fbrk
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   204
                  (mk_term_of_def gr2 "fun " descr') @ [Pretty.str ";"])) ^ "\n\n"
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   205
              else "") ^
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   206
             (if "test" mem !mode then
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   207
                Pretty.string_of (Pretty.blk (0, separate Pretty.fbrk
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   208
                  (mk_gen_of_def gr2 "fun " descr') @ [Pretty.str ";"])) ^ "\n\n"
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   209
              else ""))) gr2
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   210
         end
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   211
  end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   212
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   213
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   214
(**** case expressions ****)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   215
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   216
fun pretty_case thy defs gr dep module brack constrs (c as Const (_, T)) ts =
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   217
  let val i = length constrs
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   218
  in if length ts <= i then
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   219
       invoke_codegen thy defs dep module brack (gr, eta_expand c ts (i+1))
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   220
    else
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   221
      let
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   222
        val ts1 = Library.take (i, ts);
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   223
        val t :: ts2 = Library.drop (i, ts);
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   224
        val names = foldr add_term_names
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   225
          (map (fst o fst o dest_Var) (foldr add_term_vars [] ts1)) ts1;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   226
        val (Ts, dT) = split_last (Library.take (i+1, fst (strip_type T)));
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   227
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   228
        fun pcase gr [] [] [] = ([], gr)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   229
          | pcase gr ((cname, cargs)::cs) (t::ts) (U::Us) =
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   230
              let
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   231
                val j = length cargs;
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19886
diff changeset
   232
                val xs = Name.variant_list names (replicate j "x");
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   233
                val Us' = Library.take (j, fst (strip_type U));
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   234
                val frees = map Free (xs ~~ Us');
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   235
                val (gr0, cp) = invoke_codegen thy defs dep module false
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   236
                  (gr, list_comb (Const (cname, Us' ---> dT), frees));
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   237
                val t' = Envir.beta_norm (list_comb (t, frees));
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   238
                val (gr1, p) = invoke_codegen thy defs dep module false (gr0, t');
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   239
                val (ps, gr2) = pcase gr1 cs ts Us;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   240
              in
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   241
                ([Pretty.block [cp, Pretty.str " =>", Pretty.brk 1, p]] :: ps, gr2)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   242
              end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   243
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   244
        val (ps1, gr1) = pcase gr constrs ts1 Ts;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   245
        val ps = List.concat (separate [Pretty.brk 1, Pretty.str "| "] ps1);
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   246
        val (gr2, p) = invoke_codegen thy defs dep module false (gr1, t);
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   247
        val (gr3, ps2) = foldl_map (invoke_codegen thy defs dep module true) (gr2, ts2)
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   248
      in (gr3, (if not (null ts2) andalso brack then parens else I)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   249
        (Pretty.block (separate (Pretty.brk 1)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   250
          (Pretty.block ([Pretty.str "(case ", p, Pretty.str " of",
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   251
             Pretty.brk 1] @ ps @ [Pretty.str ")"]) :: ps2))))
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   252
      end
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   253
  end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   254
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   255
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   256
(**** constructors ****)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   257
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   258
fun pretty_constr thy defs gr dep module brack args (c as Const (s, T)) ts =
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   259
  let val i = length args
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   260
  in if i > 1 andalso length ts < i then
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   261
      invoke_codegen thy defs dep module brack (gr, eta_expand c ts i)
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   262
     else
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   263
       let
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   264
         val id = mk_qual_id module (get_const_id s gr);
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   265
         val (gr', ps) = foldl_map
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   266
           (invoke_codegen thy defs dep module (i = 1)) (gr, ts);
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   267
       in (case args of
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   268
          _ :: _ :: _ => (gr', (if brack then parens else I)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   269
            (Pretty.block [Pretty.str id, Pretty.brk 1, mk_tuple ps]))
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   270
        | _ => (gr', mk_app brack (Pretty.str id) ps))
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   271
       end
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   272
  end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   273
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   274
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   275
(**** code generators for terms and types ****)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   276
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   277
fun datatype_codegen thy defs gr dep module brack t = (case strip_comb t of
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   278
   (c as Const (s, T), ts) =>
18247
b17724cae935 code generator: case expressions, improved name resolving
haftmann
parents: 18220
diff changeset
   279
       (case Library.find_first (fn (_, {index, descr, case_name, ...}) =>
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   280
         s = case_name orelse
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   281
           AList.defined (op =) ((#3 o the o AList.lookup (op =) descr) index) s)
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   282
             (Symtab.dest (DatatypePackage.get_datatypes thy)) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
   283
          NONE => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
   284
        | SOME (tname, {index, descr, ...}) =>
18247
b17724cae935 code generator: case expressions, improved name resolving
haftmann
parents: 18220
diff changeset
   285
           if is_some (get_assoc_code thy s T) then NONE else
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   286
           let val SOME (_, _, constrs) = AList.lookup (op =) descr index
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   287
           in (case (AList.lookup (op =) constrs s, strip_type T) of
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   288
               (NONE, _) => SOME (pretty_case thy defs gr dep module brack
18247
b17724cae935 code generator: case expressions, improved name resolving
haftmann
parents: 18220
diff changeset
   289
                 ((#3 o the o AList.lookup (op =) descr) index) c ts)
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   290
             | (SOME args, (_, Type _)) => SOME (pretty_constr thy defs
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   291
                 (fst (invoke_tycodegen thy defs dep module false
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   292
                    (gr, snd (strip_type T))))
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   293
                 dep module brack args c ts)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
   294
             | _ => NONE)
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   295
           end)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
   296
 |  _ => NONE);
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   297
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   298
fun datatype_tycodegen thy defs gr dep module brack (Type (s, Ts)) =
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17261
diff changeset
   299
      (case Symtab.lookup (DatatypePackage.get_datatypes thy) s of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
   300
         NONE => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
   301
       | SOME {descr, ...} =>
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   302
           if isSome (get_assoc_type thy s) then NONE else
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   303
           let
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   304
             val (gr', ps) = foldl_map
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   305
               (invoke_tycodegen thy defs dep module false) (gr, Ts);
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   306
             val gr'' = add_dt_defs thy defs dep module gr' descr
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   307
           in SOME (gr'',
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   308
             Pretty.block ((if null Ts then [] else
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   309
               [mk_tuple ps, Pretty.str " "]) @
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   310
               [Pretty.str (mk_qual_id module (get_type_id s gr''))]))
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   311
           end)
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   312
  | datatype_tycodegen _ _ _ _ _ _ _ = NONE;
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   313
18247
b17724cae935 code generator: case expressions, improved name resolving
haftmann
parents: 18220
diff changeset
   314
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   315
(** code 2nd generation **)
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   316
20177
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   317
fun datatypes_dependency thy =
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   318
  let
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   319
    val dtnames = DatatypePackage.get_datatypes thy;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   320
    fun add_node (dtname, _) =
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   321
      let
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   322
        fun add_tycos (Type (tyco, tys)) = insert (op =) tyco #> fold add_tycos tys
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   323
          | add_tycos _ = I;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   324
        val deps = (filter (Symtab.defined dtnames) o maps (fn ty =>
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   325
          add_tycos ty [])
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   326
            o maps snd o snd o the o DatatypePackage.get_datatype_spec thy) dtname
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   327
      in
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   328
        Graph.default_node (dtname, ())
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   329
        #> fold (fn dtname' =>
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   330
             Graph.default_node (dtname', ())
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   331
             #> Graph.add_edge (dtname', dtname)
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   332
           ) deps
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   333
      end
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   334
  in
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   335
    Graph.empty
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   336
    |> Symtab.fold add_node dtnames
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   337
    |> Graph.strong_conn
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   338
  end;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   339
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   340
fun get_datatype_mut_specs thy (tycos as tyco :: _) =
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   341
  let
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   342
    val tycos' = (map (#1 o snd) o #descr o DatatypePackage.the_datatype thy) tyco;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   343
    val _ = if gen_subset (op =) (tycos, tycos') then () else
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   344
      error ("datatype constructors are not mutually recursive: " ^ (commas o map quote) tycos);
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   345
    val (vs::_, css) = split_list (map (the o DatatypePackage.get_datatype_spec thy) tycos);
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   346
  in (vs, tycos ~~ css) end;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   347
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   348
fun get_datatype_arities thy tycos sort =
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   349
  let
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   350
    val algebra = Sign.classes_of thy;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   351
    val (vs_proto, css_proto) = get_datatype_mut_specs thy tycos;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   352
    val vs = map (fn (v, vsort) => (v, Sorts.inter_sort algebra (vsort, sort))) vs_proto;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   353
    fun inst_type tyco (c, tys) =
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   354
      let
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   355
        val tys' = (map o map_atyps)
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   356
          (fn TFree (v, _) => TFree (v, the (AList.lookup (op =) vs v))) tys
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   357
      in (c, tys') end;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   358
    val css = map (fn (tyco, cs) => (tyco, (map (inst_type tyco) cs))) css_proto;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   359
    fun mk_arity tyco =
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   360
      ((tyco, map snd vs), sort);
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   361
    fun typ_of_sort ty =
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   362
      let
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   363
        val arities = map (fn (tyco, _) => ((tyco, map snd vs), sort)) css;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   364
      in ClassPackage.assume_arities_of_sort thy arities (ty, sort) end;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   365
    fun mk_cons tyco (c, tys) =
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   366
      let
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   367
        val ts = Name.give_names Name.context "a" tys;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   368
        val ty = tys ---> Type (tyco, map TFree vs);
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   369
      in list_comb (Const (c, ty), map Free ts) end;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   370
  in if forall (fn (_, cs) => forall (fn (_, tys) => forall typ_of_sort tys) cs) css
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   371
    then SOME (
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   372
      map (fn (tyco, cs) => (tyco, (mk_arity tyco, map (mk_cons tyco) cs))) css
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   373
    ) else NONE
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   374
  end;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   375
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   376
fun datatype_prove_arities tac tycos sort f thy =
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   377
  case get_datatype_arities thy tycos sort
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   378
   of NONE => thy
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   379
    | SOME insts => let
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   380
        fun proven ((tyco, asorts), sort) =
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   381
          Sorts.of_sort (Sign.classes_of thy)
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   382
            (Type (tyco, map TFree (Name.give_names Name.context "'a" asorts)), sort);
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   383
        val (arities, css) = (split_list o map_filter
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   384
          (fn (tyco, (arity, cs)) => if proven arity
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   385
            then SOME (arity, (tyco, cs)) else NONE)) insts;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   386
      in
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   387
        thy
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   388
        |> ClassPackage.prove_instance_arity tac
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   389
             arities ("", []) (f css)
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   390
      end;
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   391
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   392
fun dtyp_of_case_const thy c =
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   393
  get_first (fn (dtco, { case_name, ... }) => if case_name = c then SOME dtco else NONE)
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   394
    ((Symtab.dest o DatatypePackage.get_datatypes) thy);
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   395
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   396
fun dest_case_app cs ts tys =
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   397
  let
20177
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   398
    val abs = Name.give_names Name.context "a" (Library.drop (length ts, tys));
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   399
    val (ts', t) = split_last (ts @ map Free abs);
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   400
    val (tys', sty) = split_last tys;
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   401
    fun freenames_of t = fold_aterms
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   402
      (fn Free (v, _) => insert (op =) v | _ => I) t [];
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   403
    fun dest_case ((c, tys_decl), ty) t =
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   404
      let
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   405
        val (vs, t') = Term.strip_abs_eta (length tys_decl) t;
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   406
        val c' = list_comb (Const (c, map snd vs ---> sty), map Free vs);
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   407
      in (c', t') end;
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   408
  in (abs, ((t, sty), map2 dest_case (cs ~~ tys') ts')) end;
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   409
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   410
fun dest_case_expr thy t =
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   411
  case strip_comb t
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   412
   of (Const (c, ty), ts) =>
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   413
        (case dtyp_of_case_const thy c
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   414
         of SOME dtco =>
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   415
              let val (vs, cs) = (the o DatatypePackage.get_datatype_spec thy) dtco;
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   416
              in SOME (dest_case_app cs ts (Library.take (length cs + 1, (fst o strip_type) ty))) end
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   417
          | _ => NONE)
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   418
    | _ => NONE;
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   419
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   420
fun datatype_tac thy dtco =
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   421
  let
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   422
    val ctxt = Context.init_proof thy;
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   423
    val inject = (#inject o DatatypePackage.the_datatype thy) dtco;
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   424
    val simpset = Simplifier.context ctxt
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   425
      (MetaSimplifier.empty_ss addsimprocs [distinct_simproc]);
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   426
  in
19458
a70f1b0f09cd more precise tactics
haftmann
parents: 19346
diff changeset
   427
    (TRY o ALLGOALS o match_tac) [HOL.eq_reflection]
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   428
    THEN (
19458
a70f1b0f09cd more precise tactics
haftmann
parents: 19346
diff changeset
   429
      (ALLGOALS o match_tac) (eqTrueI :: inject)
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   430
      ORELSE (ALLGOALS o simp_tac) simpset
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   431
    )
19458
a70f1b0f09cd more precise tactics
haftmann
parents: 19346
diff changeset
   432
    THEN (ALLGOALS o match_tac) [HOL.refl, Drule.reflexive_thm]
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   433
  end;
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   434
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   435
fun get_datatype_spec_thms thy dtco =
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   436
  case DatatypePackage.get_datatype_spec thy dtco
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   437
   of SOME vs_cos =>
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   438
        SOME (vs_cos, datatype_tac thy dtco)
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   439
    | NONE => NONE;
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   440
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   441
fun get_all_datatype_cons thy =
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   442
  Symtab.fold (fn (dtco, _) => fold
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   443
    (fn (co, _) => cons (co, dtco))
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   444
      ((snd o the oo DatatypePackage.get_datatype_spec) thy dtco))
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   445
        (DatatypePackage.get_datatypes thy) [];
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   446
19599
a5c7eb37d14f added DatatypeHooks
haftmann
parents: 19458
diff changeset
   447
fun add_datatype_case_const dtco thy =
a5c7eb37d14f added DatatypeHooks
haftmann
parents: 19458
diff changeset
   448
  let
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   449
    val {case_name, index, descr, ...} = DatatypePackage.the_datatype thy dtco;
19599
a5c7eb37d14f added DatatypeHooks
haftmann
parents: 19458
diff changeset
   450
  in
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   451
    CodegenPackage.add_appconst_i (case_name, CodegenPackage.appgen_case dest_case_expr) thy
19599
a5c7eb37d14f added DatatypeHooks
haftmann
parents: 19458
diff changeset
   452
  end;
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   453
19818
5c5c1208a3fa adding case theorems for code generator
haftmann
parents: 19599
diff changeset
   454
fun add_datatype_case_defs dtco thy =
5c5c1208a3fa adding case theorems for code generator
haftmann
parents: 19599
diff changeset
   455
  let
5c5c1208a3fa adding case theorems for code generator
haftmann
parents: 19599
diff changeset
   456
    val {case_rewrites, ...} = DatatypePackage.the_datatype thy dtco
5c5c1208a3fa adding case theorems for code generator
haftmann
parents: 19599
diff changeset
   457
  in
5c5c1208a3fa adding case theorems for code generator
haftmann
parents: 19599
diff changeset
   458
    fold CodegenTheorems.add_fun case_rewrites thy
5c5c1208a3fa adding case theorems for code generator
haftmann
parents: 19599
diff changeset
   459
  end;
5c5c1208a3fa adding case theorems for code generator
haftmann
parents: 19599
diff changeset
   460
19008
14c1b2f5dda4 improved code generator devarification
haftmann
parents: 18963
diff changeset
   461
val setup = 
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18702
diff changeset
   462
  add_codegen "datatype" datatype_codegen #>
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18702
diff changeset
   463
  add_tycodegen "datatype" datatype_tycodegen #>
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   464
  CodegenTheorems.add_datatype_extr
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   465
    get_datatype_spec_thms #>
18451
5ff0244e25e8 slight clean ups
haftmann
parents: 18386
diff changeset
   466
  CodegenPackage.set_get_all_datatype_cons
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   467
    get_all_datatype_cons #>
20177
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   468
  DatatypeHooks.add (fold add_datatype_case_const) #>
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   469
  DatatypeHooks.add (fold add_datatype_case_defs)
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   470
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   471
end;