src/HOL/Tools/datatype_codegen.ML
author haftmann
Fri, 10 Aug 2007 17:04:34 +0200
changeset 24219 e558fe311376
parent 24137 8d7896398147
child 24423 ae9cd0e92423
permissions -rw-r--r--
new structure for code generator modules
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
20855
9f60d493c8fe clarified header comments
haftmann
parents: 20835
diff changeset
     5
Code generator for inductive datatypes.
12445
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
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
    10
  val get_eq: theory -> string -> thm list
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
    11
  val get_eq_datatype: theory -> string -> thm list
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
    12
  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
    13
    -> ((string * typ) list * ((term * typ) * (term * term) list)) option
23811
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
    14
  val get_case_cert: theory -> string -> thm
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
    15
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
    16
  type hook = (string * (bool * ((string * sort) list * (string * typ list) list))) list
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
    17
    -> theory -> theory
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
    18
  val codetype_hook: hook
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
    19
  val eq_hook: hook
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
    20
  val codetypes_dependency: theory -> (string * bool) list list
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
    21
  val add_codetypes_hook_bootstrap: hook -> theory -> theory
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
    22
  val the_codetypes_mut_specs: theory -> (string * bool) list
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
    23
    -> ((string * sort) list * (string * (bool * (string * typ list) list)) list)
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
    24
  val get_codetypes_arities: theory -> (string * bool) list -> sort
22331
7df6bc8cf0b0 unified arity parser/arguments;
wenzelm
parents: 22296
diff changeset
    25
    -> (string * (arity * term list)) list option
21924
fe474e69e603 simplified class_package
haftmann
parents: 21708
diff changeset
    26
  val prove_codetypes_arities: tactic -> (string * bool) list -> sort
22331
7df6bc8cf0b0 unified arity parser/arguments;
wenzelm
parents: 22296
diff changeset
    27
    -> (arity list -> (string * term list) list -> theory
22423
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
    28
      -> ((bstring * Attrib.src list) * term) list * theory)
22331
7df6bc8cf0b0 unified arity parser/arguments;
wenzelm
parents: 22296
diff changeset
    29
    -> (arity list -> (string * term list) list -> theory -> theory)
20835
haftmann
parents: 20681
diff changeset
    30
    -> theory -> theory
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
    31
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18702
diff changeset
    32
  val setup: theory -> theory
20835
haftmann
parents: 20681
diff changeset
    33
  val setup_hooks: theory -> theory
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    34
end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    35
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    36
structure DatatypeCodegen : DATATYPE_CODEGEN =
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    37
struct
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    38
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    39
open Codegen;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    40
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    41
fun mk_tuple [p] = p
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    42
  | mk_tuple ps = Pretty.block (Pretty.str "(" ::
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    43
      List.concat (separate [Pretty.str ",", Pretty.brk 1] (map single ps)) @
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    44
        [Pretty.str ")"]);
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    45
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    46
(**** datatype definition ****)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    47
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    48
(* find shortest path to constructor with no recursive arguments *)
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    49
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    50
fun find_nonempty (descr: DatatypeAux.descr) is i =
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    51
  let
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
    52
    val (_, _, constrs) = valOf (AList.lookup (op =) descr i);
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    53
    fun arg_nonempty (_, DatatypeAux.DtRec i) = if i mem is then NONE
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    54
          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
    55
      | arg_nonempty _ = SOME 0;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    56
    fun max xs = Library.foldl
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    57
      (fn (NONE, _) => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    58
        | (SOME i, SOME j) => SOME (Int.max (i, j))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    59
        | (_, NONE) => NONE) (SOME 0, xs);
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    60
    val xs = sort (int_ord o pairself snd)
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    61
      (List.mapPartial (fn (s, dts) => Option.map (pair s)
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    62
        (max (map (arg_nonempty o DatatypeAux.strip_dtyp) dts))) constrs)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    63
  in case xs of [] => NONE | x :: _ => SOME x end;
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    64
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    65
fun add_dt_defs thy defs dep module gr (descr: DatatypeAux.descr) =
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    66
  let
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    67
    val descr' = List.filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr;
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    68
    val rtnames = map (#1 o snd) (List.filter (fn (_, (_, _, cs)) =>
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    69
      exists (exists DatatypeAux.is_rec_type o snd) cs) descr');
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    70
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    71
    val (_, (tname, _, _)) :: _ = descr';
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    72
    val node_id = tname ^ " (type)";
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    73
    val module' = if_library (thyname_of_type tname thy) module;
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    74
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    75
    fun mk_dtdef gr prfx [] = (gr, [])
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    76
      | mk_dtdef gr prfx ((_, (tname, dts, cs))::xs) =
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    77
          let
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    78
            val tvs = map DatatypeAux.dest_DtTFree dts;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    79
            val sorts = map (rpair []) tvs;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    80
            val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs;
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    81
            val (gr', (_, type_id)) = mk_type_id module' tname gr;
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    82
            val (gr'', ps) =
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    83
              foldl_map (fn (gr, (cname, cargs)) =>
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    84
                foldl_map (invoke_tycodegen thy defs node_id module' false)
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    85
                  (gr, cargs) |>>>
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    86
                mk_const_id module' cname) (gr', cs');
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    87
            val (gr''', rest) = mk_dtdef gr'' "and " xs
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    88
          in
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    89
            (gr''',
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    90
             Pretty.block (Pretty.str prfx ::
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    91
               (if null tvs then [] else
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    92
                  [mk_tuple (map Pretty.str tvs), Pretty.str " "]) @
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    93
               [Pretty.str (type_id ^ " ="), Pretty.brk 1] @
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    94
               List.concat (separate [Pretty.brk 1, Pretty.str "| "]
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    95
                 (map (fn (ps', (_, cname)) => [Pretty.block
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    96
                   (Pretty.str cname ::
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    97
                    (if null ps' then [] else
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    98
                     List.concat ([Pretty.str " of", Pretty.brk 1] ::
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    99
                       separate [Pretty.str " *", Pretty.brk 1]
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   100
                         (map single ps'))))]) ps))) :: rest)
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   101
          end;
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   102
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   103
    fun mk_term_of_def gr prfx [] = []
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   104
      | 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
   105
          let
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   106
            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
   107
            val sorts = map (rpair []) tvs;
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   108
            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
   109
            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
   110
            val T = Type (tname, dts');
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   111
            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
   112
            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
   113
              let val args = map (fn i =>
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   114
                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
   115
              in ("  | ", Pretty.blk (4,
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   116
                [Pretty.str prfx, mk_term_of gr module' false T, Pretty.brk 1,
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   117
                 if null Ts then Pretty.str (snd (get_const_id cname gr))
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   118
                 else parens (Pretty.block
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   119
                   [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
   120
                    Pretty.brk 1, mk_tuple args]),
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   121
                 Pretty.str " =", Pretty.brk 1] @
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   122
                 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
   123
                   ([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
   124
                     mk_type false (Ts ---> T), Pretty.str ")"] ::
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   125
                    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
   126
                      Pretty.brk 1, x]]) (args ~~ Ts)))))
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   127
              end) (prfx, cs')
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   128
          in eqs @ rest end;
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   129
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   130
    fun mk_gen_of_def gr prfx [] = []
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   131
      | mk_gen_of_def gr prfx ((i, (tname, dts, cs)) :: xs) =
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   132
          let
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   133
            val tvs = map DatatypeAux.dest_DtTFree dts;
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   134
            val sorts = map (rpair []) tvs;
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   135
            val (cs1, cs2) =
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   136
              List.partition (exists DatatypeAux.is_rec_type o snd) cs;
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
   137
            val SOME (cname, _) = find_nonempty descr [i] i;
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   138
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   139
            fun mk_delay p = Pretty.block
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   140
              [Pretty.str "fn () =>", Pretty.brk 1, p];
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   141
15397
5260ac75e07c Fixed bug in mk_gen_of_def that could cause non-termination of the generator
berghofe
parents: 14981
diff changeset
   142
            fun mk_constr s b (cname, dts) =
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   143
              let
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   144
                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
   145
                    (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
   146
                  [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
   147
                     else "j")]) dts;
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   148
                val (_, id) = get_const_id cname gr
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   149
              in case gs of
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   150
                  _ :: _ :: _ => Pretty.block
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   151
                    [Pretty.str id, Pretty.brk 1, mk_tuple gs]
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   152
                | _ => mk_app false (Pretty.str id) (map parens gs)
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   153
              end;
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   154
15397
5260ac75e07c Fixed bug in mk_gen_of_def that could cause non-termination of the generator
berghofe
parents: 14981
diff changeset
   155
            fun mk_choice [c] = mk_constr "(i-1)" false c
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   156
              | mk_choice cs = Pretty.block [Pretty.str "one_of",
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   157
                  Pretty.brk 1, Pretty.blk (1, Pretty.str "[" ::
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   158
                  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
   159
                    (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
   160
                  [Pretty.str "]"]), Pretty.brk 1, Pretty.str "()"];
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   161
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   162
            val gs = map (Pretty.str o suffix "G" o strip_tname) tvs;
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   163
            val gen_name = "gen_" ^ snd (get_type_id tname gr)
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   164
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   165
          in
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   166
            Pretty.blk (4, separate (Pretty.brk 1) 
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   167
                (Pretty.str (prfx ^ gen_name ^
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   168
                   (if null cs1 then "" else "'")) :: gs @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   169
                 (if null cs1 then [] else [Pretty.str "i"]) @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   170
                 [Pretty.str "j"]) @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   171
              [Pretty.str " =", Pretty.brk 1] @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   172
              (if not (null cs1) andalso not (null cs2)
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   173
               then [Pretty.str "frequency", Pretty.brk 1,
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   174
                 Pretty.blk (1, [Pretty.str "[",
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   175
                   mk_tuple [Pretty.str "i", mk_delay (mk_choice cs1)],
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   176
                   Pretty.str ",", Pretty.fbrk,
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   177
                   mk_tuple [Pretty.str "1", mk_delay (mk_choice cs2)],
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   178
                   Pretty.str "]"]), Pretty.brk 1, Pretty.str "()"]
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   179
               else if null cs2 then
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   180
                 [Pretty.block [Pretty.str "(case", Pretty.brk 1,
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   181
                   Pretty.str "i", Pretty.brk 1, Pretty.str "of",
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   182
                   Pretty.brk 1, Pretty.str "0 =>", Pretty.brk 1,
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   183
                   mk_constr "0" true (cname, valOf (AList.lookup (op =) cs cname)),
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   184
                   Pretty.brk 1, Pretty.str "| _ =>", Pretty.brk 1,
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   185
                   mk_choice cs1, Pretty.str ")"]]
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   186
               else [mk_choice cs2])) ::
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   187
            (if null cs1 then []
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   188
             else [Pretty.blk (4, separate (Pretty.brk 1) 
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   189
                 (Pretty.str ("and " ^ gen_name) :: gs @ [Pretty.str "i"]) @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   190
               [Pretty.str " =", Pretty.brk 1] @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   191
               separate (Pretty.brk 1) (Pretty.str (gen_name ^ "'") :: gs @
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   192
                 [Pretty.str "i", Pretty.str "i"]))]) @
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   193
            mk_gen_of_def gr "and " xs
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   194
          end
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   195
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   196
  in
20681
0e4df994ad34 Fixed bug concerning the generation of identifiers for
berghofe
parents: 20608
diff changeset
   197
    ((add_edge_acyclic (node_id, dep) gr
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   198
        handle Graph.CYCLES _ => gr) handle Graph.UNDEF _ =>
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   199
         let
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   200
           val gr1 = add_edge (node_id, dep)
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   201
             (new_node (node_id, (NONE, "", "")) gr);
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   202
           val (gr2, dtdef) = mk_dtdef gr1 "datatype " descr';
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   203
         in
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   204
           map_node node_id (K (NONE, module',
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   205
             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
   206
               [Pretty.str ";"])) ^ "\n\n" ^
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   207
             (if "term_of" mem !mode then
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   208
                Pretty.string_of (Pretty.blk (0, separate Pretty.fbrk
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   209
                  (mk_term_of_def gr2 "fun " descr') @ [Pretty.str ";"])) ^ "\n\n"
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   210
              else "") ^
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   211
             (if "test" mem !mode then
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   212
                Pretty.string_of (Pretty.blk (0, separate Pretty.fbrk
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   213
                  (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
   214
              else ""))) gr2
20681
0e4df994ad34 Fixed bug concerning the generation of identifiers for
berghofe
parents: 20608
diff changeset
   215
         end,
0e4df994ad34 Fixed bug concerning the generation of identifiers for
berghofe
parents: 20608
diff changeset
   216
     module')
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   217
  end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   218
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   219
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   220
(**** case expressions ****)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   221
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   222
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
   223
  let val i = length constrs
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   224
  in if length ts <= i then
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   225
       invoke_codegen thy defs dep module brack (gr, eta_expand c ts (i+1))
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   226
    else
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   227
      let
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   228
        val ts1 = Library.take (i, ts);
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   229
        val t :: ts2 = Library.drop (i, ts);
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   230
        val names = foldr add_term_names
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   231
          (map (fst o fst o dest_Var) (foldr add_term_vars [] ts1)) ts1;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   232
        val (Ts, dT) = split_last (Library.take (i+1, fst (strip_type T)));
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   233
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   234
        fun pcase gr [] [] [] = ([], gr)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   235
          | pcase gr ((cname, cargs)::cs) (t::ts) (U::Us) =
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   236
              let
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   237
                val j = length cargs;
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19886
diff changeset
   238
                val xs = Name.variant_list names (replicate j "x");
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   239
                val Us' = Library.take (j, fst (strip_type U));
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   240
                val frees = map Free (xs ~~ Us');
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   241
                val (gr0, cp) = invoke_codegen thy defs dep module false
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   242
                  (gr, list_comb (Const (cname, Us' ---> dT), frees));
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   243
                val t' = Envir.beta_norm (list_comb (t, frees));
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   244
                val (gr1, p) = invoke_codegen thy defs dep module false (gr0, t');
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   245
                val (ps, gr2) = pcase gr1 cs ts Us;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   246
              in
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   247
                ([Pretty.block [cp, Pretty.str " =>", Pretty.brk 1, p]] :: ps, gr2)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   248
              end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   249
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   250
        val (ps1, gr1) = pcase gr constrs ts1 Ts;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   251
        val ps = List.concat (separate [Pretty.brk 1, Pretty.str "| "] ps1);
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   252
        val (gr2, p) = invoke_codegen thy defs dep module false (gr1, t);
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   253
        val (gr3, ps2) = foldl_map (invoke_codegen thy defs dep module true) (gr2, ts2)
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   254
      in (gr3, (if not (null ts2) andalso brack then parens else I)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   255
        (Pretty.block (separate (Pretty.brk 1)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   256
          (Pretty.block ([Pretty.str "(case ", p, Pretty.str " of",
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   257
             Pretty.brk 1] @ ps @ [Pretty.str ")"]) :: ps2))))
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   258
      end
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   259
  end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   260
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   261
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   262
(**** constructors ****)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   263
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   264
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
   265
  let val i = length args
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   266
  in if i > 1 andalso length ts < i then
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   267
      invoke_codegen thy defs dep module brack (gr, eta_expand c ts i)
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   268
     else
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   269
       let
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   270
         val id = mk_qual_id module (get_const_id s gr);
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   271
         val (gr', ps) = foldl_map
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   272
           (invoke_codegen thy defs dep module (i = 1)) (gr, ts);
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   273
       in (case args of
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   274
          _ :: _ :: _ => (gr', (if brack then parens else I)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   275
            (Pretty.block [Pretty.str id, Pretty.brk 1, mk_tuple ps]))
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   276
        | _ => (gr', mk_app brack (Pretty.str id) ps))
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   277
       end
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   278
  end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   279
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   280
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   281
(**** code generators for terms and types ****)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   282
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   283
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
   284
   (c as Const (s, T), ts) =>
22778
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   285
     (case DatatypePackage.datatype_of_case thy s of
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   286
        SOME {index, descr, ...} =>
22921
475ff421a6a3 consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents: 22778
diff changeset
   287
          if is_some (get_assoc_code thy (s, T)) then NONE else
22778
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   288
          SOME (pretty_case thy defs gr dep module brack
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   289
            (#3 (the (AList.lookup op = descr index))) c ts)
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   290
      | NONE => case (DatatypePackage.datatype_of_constr thy s, strip_type T) of
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   291
        (SOME {index, descr, ...}, (_, U as Type _)) =>
22921
475ff421a6a3 consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents: 22778
diff changeset
   292
          if is_some (get_assoc_code thy (s, T)) then NONE else
22778
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   293
          let val SOME args = AList.lookup op =
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   294
            (#3 (the (AList.lookup op = descr index))) s
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   295
          in
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   296
            SOME (pretty_constr thy defs
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   297
              (fst (invoke_tycodegen thy defs dep module false (gr, U)))
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   298
              dep module brack args c ts)
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   299
          end
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   300
      | _ => NONE)
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   301
 | _ => NONE);
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   302
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   303
fun datatype_tycodegen thy defs gr dep module brack (Type (s, Ts)) =
22778
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   304
      (case DatatypePackage.get_datatype thy s of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
   305
         NONE => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
   306
       | SOME {descr, ...} =>
22921
475ff421a6a3 consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents: 22778
diff changeset
   307
           if is_some (get_assoc_type thy s) then NONE else
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   308
           let
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   309
             val (gr', ps) = foldl_map
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   310
               (invoke_tycodegen thy defs dep module false) (gr, Ts);
20681
0e4df994ad34 Fixed bug concerning the generation of identifiers for
berghofe
parents: 20608
diff changeset
   311
             val (gr'', module') = add_dt_defs thy defs dep module gr' descr;
0e4df994ad34 Fixed bug concerning the generation of identifiers for
berghofe
parents: 20608
diff changeset
   312
             val (gr''', tyid) = mk_type_id module' s gr''
0e4df994ad34 Fixed bug concerning the generation of identifiers for
berghofe
parents: 20608
diff changeset
   313
           in SOME (gr''',
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   314
             Pretty.block ((if null Ts then [] else
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   315
               [mk_tuple ps, Pretty.str " "]) @
20681
0e4df994ad34 Fixed bug concerning the generation of identifiers for
berghofe
parents: 20608
diff changeset
   316
               [Pretty.str (mk_qual_id module tyid)]))
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   317
           end)
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   318
  | datatype_tycodegen _ _ _ _ _ _ _ = NONE;
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   319
18247
b17724cae935 code generator: case expressions, improved name resolving
haftmann
parents: 18220
diff changeset
   320
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   321
(** datatypes for code 2nd generation **)
20177
0af885e3dabf hooks now take string list as arguments (mutual datatypes); some nice combinators in datatype_codegen
haftmann
parents: 20105
diff changeset
   322
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   323
fun dtyp_of_case_const thy c =
22778
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   324
  Option.map (fn {descr, index, ...} => #1 (the (AList.lookup op = descr index)))
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   325
    (DatatypePackage.datatype_of_case thy c);
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   326
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   327
fun dest_case_app cs ts tys =
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   328
  let
22052
792c18b8f214 improved case patterns
haftmann
parents: 22051
diff changeset
   329
    val names = (Name.make_context o map fst) (fold Term.add_tfrees ts []);
792c18b8f214 improved case patterns
haftmann
parents: 22051
diff changeset
   330
    val abs = Name.names names "a" (Library.drop (length ts, tys));
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   331
    val (ts', t) = split_last (ts @ map Free abs);
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   332
    val (tys', sty) = split_last tys;
22051
2b8909d9d66a added undefined in cases
haftmann
parents: 22047
diff changeset
   333
    fun dest_case ((c, tys_decl), ty) t =
2b8909d9d66a added undefined in cases
haftmann
parents: 22047
diff changeset
   334
      let
2b8909d9d66a added undefined in cases
haftmann
parents: 22047
diff changeset
   335
        val (vs, t') = Term.strip_abs_eta (length tys_decl) t;
2b8909d9d66a added undefined in cases
haftmann
parents: 22047
diff changeset
   336
        val c' = list_comb (Const (c, map snd vs ---> sty), map Free vs);
2b8909d9d66a added undefined in cases
haftmann
parents: 22047
diff changeset
   337
      in case t'
22480
b20bc8029edb switched exception from arbitrary to undefined
haftmann
parents: 22435
diff changeset
   338
       of Const ("HOL.undefined", _) => NONE
22051
2b8909d9d66a added undefined in cases
haftmann
parents: 22047
diff changeset
   339
        | _ => SOME (c', t')
2b8909d9d66a added undefined in cases
haftmann
parents: 22047
diff changeset
   340
      end;
22047
ff91fd74bb71 handling for "undefined" in case expressions
haftmann
parents: 21924
diff changeset
   341
  in (abs, ((t, sty), map2 dest_case (cs ~~ tys') ts' |> map_filter I)) end;
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   342
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   343
fun dest_case_expr thy t =
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   344
  case strip_comb t
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   345
   of (Const (c, ty), ts) =>
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   346
        (case dtyp_of_case_const thy c
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   347
         of SOME dtco =>
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   348
              let val (vs, cs) = (the o DatatypePackage.get_datatype_spec thy) dtco;
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   349
              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
   350
          | _ => NONE)
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   351
    | _ => NONE;
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   352
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   353
fun mk_distinct cos =
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   354
  let
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   355
    fun sym_product [] = []
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   356
      | sym_product (x::xs) = map (pair x) xs @ sym_product xs;
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   357
    fun mk_co_args (co, tys) ctxt =
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   358
      let
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   359
        val names = Name.invents ctxt "a" (length tys);
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   360
        val ctxt' = fold Name.declare names ctxt;
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   361
        val vs = map2 (curry Free) names tys;
20835
haftmann
parents: 20681
diff changeset
   362
      in (vs, ctxt') end;
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   363
    fun mk_dist ((co1, tys1), (co2, tys2)) =
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   364
      let
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   365
        val ((xs1, xs2), _) = Name.context
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   366
          |> mk_co_args (co1, tys1)
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   367
          ||>> mk_co_args (co2, tys2);
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   368
        val prem = HOLogic.mk_eq
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   369
          (list_comb (co1, xs1), list_comb (co2, xs2));
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   370
        val t = HOLogic.mk_not prem;
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   371
      in HOLogic.mk_Trueprop t end;
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   372
  in map mk_dist (sym_product cos) end;
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   373
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   374
local
20835
haftmann
parents: 20681
diff changeset
   375
  val not_sym = thm "HOL.not_sym";
haftmann
parents: 20681
diff changeset
   376
  val not_false_true = iffD2 OF [nth (thms "HOL.simp_thms") 7, TrueI];
21546
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   377
  val refl = thm "refl";
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   378
  val eqTrueI = thm "eqTrueI";
21516
c2a116a2c4fd ProofContext.init;
wenzelm
parents: 21454
diff changeset
   379
in
c2a116a2c4fd ProofContext.init;
wenzelm
parents: 21454
diff changeset
   380
c2a116a2c4fd ProofContext.init;
wenzelm
parents: 21454
diff changeset
   381
fun get_eq_datatype thy dtco =
20835
haftmann
parents: 20681
diff changeset
   382
  let
haftmann
parents: 20681
diff changeset
   383
    val SOME (vs, cs) = DatatypePackage.get_datatype_spec thy dtco;
haftmann
parents: 20681
diff changeset
   384
    fun mk_triv_inject co =
haftmann
parents: 20681
diff changeset
   385
      let
haftmann
parents: 20681
diff changeset
   386
        val ct' = Thm.cterm_of thy
haftmann
parents: 20681
diff changeset
   387
          (Const (co, Type (dtco, map (fn (v, sort) => TVar ((v, 0), sort)) vs)))
haftmann
parents: 20681
diff changeset
   388
        val cty' = Thm.ctyp_of_term ct';
haftmann
parents: 20681
diff changeset
   389
        val SOME (ct, cty) = fold_aterms (fn Var (v, ty) =>
haftmann
parents: 20681
diff changeset
   390
          (K o SOME) (Thm.cterm_of thy (Var (v, Thm.typ_of cty')), Thm.ctyp_of thy ty) | _ => I)
21546
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   391
          (Thm.prop_of refl) NONE;
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   392
      in eqTrueI OF [Thm.instantiate ([(cty, cty')], [(ct, ct')]) refl] end;
20835
haftmann
parents: 20681
diff changeset
   393
    val inject1 = map_filter (fn (co, []) => SOME (mk_triv_inject co) | _ => NONE) cs
haftmann
parents: 20681
diff changeset
   394
    val inject2 = (#inject o DatatypePackage.the_datatype thy) dtco;
21516
c2a116a2c4fd ProofContext.init;
wenzelm
parents: 21454
diff changeset
   395
    val ctxt = ProofContext.init thy;
20835
haftmann
parents: 20681
diff changeset
   396
    val simpset = Simplifier.context ctxt
haftmann
parents: 20681
diff changeset
   397
      (MetaSimplifier.empty_ss addsimprocs [distinct_simproc]);
haftmann
parents: 20681
diff changeset
   398
    val cos = map (fn (co, tys) =>
haftmann
parents: 20681
diff changeset
   399
        (Const (co, tys ---> Type (dtco, map TFree vs)), tys)) cs;
haftmann
parents: 20681
diff changeset
   400
    val tac = ALLGOALS (simp_tac simpset)
haftmann
parents: 20681
diff changeset
   401
      THEN ALLGOALS (ProofContext.fact_tac [not_false_true, TrueI]);
haftmann
parents: 20681
diff changeset
   402
    val distinct =
haftmann
parents: 20681
diff changeset
   403
      mk_distinct cos
haftmann
parents: 20681
diff changeset
   404
      |> map (fn t => Goal.prove_global thy [] [] t (K tac))
haftmann
parents: 20681
diff changeset
   405
      |> (fn thms => thms @ map (fn thm => not_sym OF [thm]) thms)
haftmann
parents: 20681
diff changeset
   406
  in inject1 @ inject2 @ distinct end;
21516
c2a116a2c4fd ProofContext.init;
wenzelm
parents: 21454
diff changeset
   407
c2a116a2c4fd ProofContext.init;
wenzelm
parents: 21454
diff changeset
   408
end;
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   409
23811
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   410
fun get_case_cert thy tyco =
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   411
  let
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   412
    val raw_thms =
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   413
      (#case_rewrites o DatatypePackage.the_datatype thy) tyco;
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   414
    val thms as hd_thm :: _ = raw_thms
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   415
      |> Conjunction.intr_balanced
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   416
      |> Drule.unvarify
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   417
      |> Conjunction.elim_balanced (length raw_thms)
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   418
      |> map Simpdata.mk_meta_eq
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   419
      |> map Drule.zero_var_indexes
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   420
    val params = fold_aterms (fn (Free (v, _)) => insert (op =) v
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   421
      | _ => I) (Thm.prop_of hd_thm) [];
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   422
    val rhs = hd_thm
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   423
      |> Thm.prop_of
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   424
      |> Logic.dest_equals
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   425
      |> fst
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   426
      |> Term.strip_comb
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   427
      |> apsnd (fst o split_last)
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   428
      |> list_comb;
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   429
    val lhs = Free (Name.variant params "case", Term.fastype_of rhs);
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   430
    val asm = (Thm.cterm_of thy o Logic.mk_equals) (lhs, rhs);
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   431
  in
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   432
    thms
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   433
    |> Conjunction.intr_balanced
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   434
    |> MetaSimplifier.rewrite_rule [(Thm.symmetric o Thm.assume) asm]
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   435
    |> Thm.implies_intr asm
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   436
    |> Thm.generalize ([], params) 0
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   437
    |> Thm.varifyT
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   438
  end;
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   439
19818
5c5c1208a3fa adding case theorems for code generator
haftmann
parents: 19599
diff changeset
   440
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   441
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   442
(** codetypes for code 2nd generation **)
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   443
20835
haftmann
parents: 20681
diff changeset
   444
(* abstraction over datatypes vs. type copies *)
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   445
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   446
fun codetypes_dependency thy =
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   447
  let
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   448
    val names =
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   449
      map (rpair true) (Symtab.keys (DatatypePackage.get_datatypes thy))
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   450
        @ map (rpair false) (TypecopyPackage.get_typecopies thy);
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   451
    fun add_node (name, is_dt) =
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   452
      let
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   453
        fun add_tycos (Type (tyco, tys)) = insert (op =) tyco #> fold add_tycos tys
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   454
          | add_tycos _ = I;
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   455
        val tys = if is_dt then
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   456
            (maps snd o snd o the o DatatypePackage.get_datatype_spec thy) name
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   457
          else
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   458
            [(#typ o the o TypecopyPackage.get_typecopy_info thy) name]
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   459
        val deps = (filter (AList.defined (op =) names) o maps (fn ty =>
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   460
          add_tycos ty [])) tys;
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   461
      in
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   462
        Graph.default_node (name, ())
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   463
        #> fold (fn name' =>
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   464
             Graph.default_node (name', ())
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   465
             #> Graph.add_edge (name', name)
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   466
           ) deps
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   467
      end
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   468
  in
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   469
    Graph.empty
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   470
    |> fold add_node names
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   471
    |> Graph.strong_conn
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   472
    |> map (AList.make (the o AList.lookup (op =) names))
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   473
  end;
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   474
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   475
fun get_spec thy (dtco, true) =
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   476
      (the o DatatypePackage.get_datatype_spec thy) dtco
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   477
  | get_spec thy (tyco, false) =
20835
haftmann
parents: 20681
diff changeset
   478
      TypecopyPackage.get_spec thy tyco;
haftmann
parents: 20681
diff changeset
   479
haftmann
parents: 20681
diff changeset
   480
local
haftmann
parents: 20681
diff changeset
   481
  fun get_eq_thms thy tyco = case DatatypePackage.get_datatype thy tyco
haftmann
parents: 20681
diff changeset
   482
   of SOME _ => get_eq_datatype thy tyco
haftmann
parents: 20681
diff changeset
   483
    | NONE => [TypecopyPackage.get_eq thy tyco];
21546
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   484
  fun constrain_op_eq_thms thy thms =
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   485
    let
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   486
      fun add_eq (Const ("op =", ty)) =
23712
haftmann
parents: 23513
diff changeset
   487
            fold (insert (eq_fst (op =))) (Term.add_tvarsT ty [])
21546
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   488
        | add_eq _ =
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   489
            I
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   490
      val eqs = fold (fold_aterms add_eq o Thm.prop_of) thms [];
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   491
      val instT = map (fn (v_i, sort) =>
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   492
        (Thm.ctyp_of thy (TVar (v_i, sort)),
23712
haftmann
parents: 23513
diff changeset
   493
           Thm.ctyp_of thy (TVar (v_i, Sorts.inter_sort (Sign.classes_of thy)
haftmann
parents: 23513
diff changeset
   494
             (sort, [HOLogic.class_eq]))))) eqs;
21546
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   495
    in
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   496
      thms
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   497
      |> map (Thm.instantiate (instT, []))
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   498
    end;
20835
haftmann
parents: 20681
diff changeset
   499
in
haftmann
parents: 20681
diff changeset
   500
  fun get_eq thy tyco =
haftmann
parents: 20681
diff changeset
   501
    get_eq_thms thy tyco
haftmann
parents: 20681
diff changeset
   502
    |> maps ((#mk o #mk_rews o snd o MetaSimplifier.rep_ss o Simplifier.simpset_of) thy)
21546
268b6bed0cc8 removed HOL structure
haftmann
parents: 21516
diff changeset
   503
    |> constrain_op_eq_thms thy
20835
haftmann
parents: 20681
diff changeset
   504
end;
haftmann
parents: 20681
diff changeset
   505
haftmann
parents: 20681
diff changeset
   506
type hook = (string * (bool * ((string * sort) list * (string * typ list) list))) list
haftmann
parents: 20681
diff changeset
   507
  -> theory -> theory;
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   508
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   509
fun add_codetypes_hook_bootstrap hook thy =
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   510
  let
20835
haftmann
parents: 20681
diff changeset
   511
    fun add_spec thy (tyco, is_dt) =
haftmann
parents: 20681
diff changeset
   512
      (tyco, (is_dt, get_spec thy (tyco, is_dt)));
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   513
    fun datatype_hook dtcos thy =
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   514
      hook (map (add_spec thy) (map (rpair true) dtcos)) thy;
20835
haftmann
parents: 20681
diff changeset
   515
    fun typecopy_hook ((tyco, _)) thy =
haftmann
parents: 20681
diff changeset
   516
      hook ([(tyco, (false, TypecopyPackage.get_spec thy tyco))]) thy;
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   517
  in
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   518
    thy
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   519
    |> fold hook ((map o map) (add_spec thy) (codetypes_dependency thy))
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   520
    |> DatatypeHooks.add datatype_hook
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   521
    |> TypecopyPackage.add_hook typecopy_hook
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   522
  end;
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   523
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   524
fun the_codetypes_mut_specs thy ([(tyco, is_dt)]) =
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   525
      let
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   526
        val (vs, cs) = get_spec thy (tyco, is_dt)
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   527
      in (vs, [(tyco, (is_dt, cs))]) end
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   528
  | the_codetypes_mut_specs thy (tycos' as (tyco, true) :: _) =
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   529
      let
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   530
        val tycos = map fst tycos';
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   531
        val tycos'' = (map (#1 o snd) o #descr o DatatypePackage.the_datatype thy) tyco;
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   532
        val _ = if gen_subset (op =) (tycos, tycos'') then () else
22175
d9e3e4c30d6b adjusted names
haftmann
parents: 22052
diff changeset
   533
          error ("type constructors are not mutually recursive: " ^ (commas o map quote) tycos);
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   534
        val (vs::_, css) = split_list (map (the o DatatypePackage.get_datatype_spec thy) tycos);
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   535
      in (vs, map2 (fn (tyco, is_dt) => fn cs => (tyco, (is_dt, cs))) tycos' css) end;
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   536
20835
haftmann
parents: 20681
diff changeset
   537
haftmann
parents: 20681
diff changeset
   538
(* registering code types in code generator *)
haftmann
parents: 20681
diff changeset
   539
22435
16e6ddc30f92 clarified code
haftmann
parents: 22429
diff changeset
   540
fun codetype_hook dtspecs =
24219
e558fe311376 new structure for code generator modules
haftmann
parents: 24137
diff changeset
   541
  fold (fn (dtco, (_, spec)) => Code.add_datatype (dtco, spec)) dtspecs;
20835
haftmann
parents: 20681
diff changeset
   542
haftmann
parents: 20681
diff changeset
   543
haftmann
parents: 20681
diff changeset
   544
(* instrumentalizing the sort algebra *)
haftmann
parents: 20681
diff changeset
   545
22423
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   546
fun get_codetypes_arities thy tycos sort =
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   547
  let
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   548
    val pp = Sign.pp thy;
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   549
    val algebra = Sign.classes_of thy;
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   550
    val (vs_proto, css_proto) = the_codetypes_mut_specs thy tycos;
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   551
    val vs = map (fn (v, vsort) => (v, Sorts.inter_sort algebra (vsort, sort))) vs_proto;
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   552
    val css = map (fn (tyco, (_, cs)) => (tyco, cs)) css_proto;
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   553
    val algebra' = algebra
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   554
      |> fold (fn (tyco, _) =>
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   555
           Sorts.add_arities pp (tyco, map (fn class => (class, map snd vs)) sort)) css;
24219
e558fe311376 new structure for code generator modules
haftmann
parents: 24137
diff changeset
   556
    fun typ_sort_inst ty = CodeUnit.typ_sort_inst algebra' (Logic.varifyT ty, sort);
22423
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   557
    val venv = Vartab.empty
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   558
      |> fold (fn (v, sort) => Vartab.update_new ((v, 0), sort)) vs
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   559
      |> fold (fn (_, cs) => fold (fn (_, tys) => fold typ_sort_inst tys) cs) css;
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   560
    fun inst (v, _) = (v, (the o Vartab.lookup venv) (v, 0));
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   561
    val vs' = map inst vs;
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   562
    fun mk_arity tyco = (tyco, map snd vs', sort);
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   563
    fun mk_cons tyco (c, tys) =
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   564
      let
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   565
        val tys' = (map o Term.map_type_tfree) (TFree o inst) tys;
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   566
        val ts = Name.names Name.context "a" tys';
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   567
        val ty = (tys' ---> Type (tyco, map TFree vs'));
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   568
      in list_comb (Const (c, ty), map Free ts) end;
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   569
  in
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   570
    map (fn (tyco, cs) => (tyco, (mk_arity tyco, map (mk_cons tyco) cs))) css |> SOME
c1836b14c63a dropped code datatype certificates
haftmann
parents: 22331
diff changeset
   571
  end handle Class_Error => NONE;
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   572
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   573
fun prove_codetypes_arities tac tycos sort f after_qed thy =
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   574
  case get_codetypes_arities thy tycos sort
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   575
   of NONE => thy
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   576
    | SOME insts => let
22331
7df6bc8cf0b0 unified arity parser/arguments;
wenzelm
parents: 22296
diff changeset
   577
        fun proven (tyco, asorts, sort) =
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   578
          Sorts.of_sort (Sign.classes_of thy)
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   579
            (Type (tyco, map TFree (Name.names Name.context "'a" asorts)), sort);
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   580
        val (arities, css) = (split_list o map_filter
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   581
          (fn (tyco, (arity, cs)) => if proven arity
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   582
            then NONE else SOME (arity, (tyco, cs)))) insts;
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   583
      in
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   584
        thy
21565
bd28361f4c5b simplified '?' operator;
wenzelm
parents: 21546
diff changeset
   585
        |> not (null arities) ? (
20835
haftmann
parents: 20681
diff changeset
   586
            f arities css
haftmann
parents: 20681
diff changeset
   587
            #-> (fn defs =>
24219
e558fe311376 new structure for code generator modules
haftmann
parents: 24137
diff changeset
   588
              Class.prove_instance_arity tac arities defs
20835
haftmann
parents: 20681
diff changeset
   589
            #> after_qed arities css))
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   590
      end;
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   591
20835
haftmann
parents: 20681
diff changeset
   592
haftmann
parents: 20681
diff changeset
   593
(* operational equality *)
haftmann
parents: 20681
diff changeset
   594
21046
fe1db2f991a7 moved HOL code generator setup to Code_Generator
haftmann
parents: 21012
diff changeset
   595
fun eq_hook specs =
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   596
  let
20835
haftmann
parents: 20681
diff changeset
   597
    fun add_eq_thms (dtco, (_, (vs, cs))) thy =
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   598
      let
24137
8d7896398147 replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
wenzelm
parents: 23811
diff changeset
   599
        val thy_ref = Theory.check_thy thy;
22554
d1499fff65d8 simplified constant representation in code generator
haftmann
parents: 22480
diff changeset
   600
        val const = ("op =", SOME dtco);
20835
haftmann
parents: 20681
diff changeset
   601
        val get_thms = (fn () => get_eq (Theory.deref thy_ref) dtco |> rev);
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   602
      in
24219
e558fe311376 new structure for code generator modules
haftmann
parents: 24137
diff changeset
   603
        Code.add_funcl (const, Susp.delay get_thms) thy
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   604
      end;
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   605
  in
24219
e558fe311376 new structure for code generator modules
haftmann
parents: 24137
diff changeset
   606
    prove_codetypes_arities (Class.intro_classes_tac [])
20835
haftmann
parents: 20681
diff changeset
   607
      (map (fn (tyco, (is_dt, _)) => (tyco, is_dt)) specs)
21046
fe1db2f991a7 moved HOL code generator setup to Code_Generator
haftmann
parents: 21012
diff changeset
   608
      [HOLogic.class_eq] ((K o K o pair) []) ((K o K) (fold add_eq_thms specs))
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   609
  end;
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   610
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   611
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   612
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   613
(** theory setup **)
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   614
23513
haftmann
parents: 22921
diff changeset
   615
fun add_datatype_case_const dtco thy =
haftmann
parents: 22921
diff changeset
   616
  let
haftmann
parents: 22921
diff changeset
   617
    val {case_name, index, descr, ...} = DatatypePackage.the_datatype thy dtco;
haftmann
parents: 22921
diff changeset
   618
  in
24219
e558fe311376 new structure for code generator modules
haftmann
parents: 24137
diff changeset
   619
    CodePackage.add_appconst (case_name, CodePackage.appgen_case dest_case_expr) thy
23513
haftmann
parents: 22921
diff changeset
   620
  end;
haftmann
parents: 22921
diff changeset
   621
haftmann
parents: 22921
diff changeset
   622
fun add_datatype_case_defs dtco thy =
haftmann
parents: 22921
diff changeset
   623
  let
haftmann
parents: 22921
diff changeset
   624
    val {case_rewrites, ...} = DatatypePackage.the_datatype thy dtco
haftmann
parents: 22921
diff changeset
   625
  in
24219
e558fe311376 new structure for code generator modules
haftmann
parents: 24137
diff changeset
   626
    fold_rev (Code.add_func true) case_rewrites thy
23513
haftmann
parents: 22921
diff changeset
   627
  end;
haftmann
parents: 22921
diff changeset
   628
19008
14c1b2f5dda4 improved code generator devarification
haftmann
parents: 18963
diff changeset
   629
val setup = 
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   630
  add_codegen "datatype" datatype_codegen
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   631
  #> add_tycodegen "datatype" datatype_tycodegen 
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   632
  #> DatatypeHooks.add (fold add_datatype_case_const)
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   633
  #> DatatypeHooks.add (fold add_datatype_case_defs)
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   634
20835
haftmann
parents: 20681
diff changeset
   635
val setup_hooks =
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   636
  add_codetypes_hook_bootstrap codetype_hook
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   637
  #> add_codetypes_hook_bootstrap eq_hook
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   638
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   639
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   640
end;