src/HOL/Tools/datatype_codegen.ML
author haftmann
Sun, 24 May 2009 15:02:21 +0200
changeset 31246 251a34663242
parent 31156 90fed3d4430f
child 31458 b1cf26f2919b
permissions -rw-r--r--
exported find_shorzes_path
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24428
diff changeset
     1
(*  Title:      HOL/Tools/datatype_codegen.ML
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 28965
diff changeset
     2
    Author:     Stefan Berghofer and Florian Haftmann, TU Muenchen
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
     3
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25505
diff changeset
     4
Code generator facilities for inductive datatypes.
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
     5
*)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
     6
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
     7
signature DATATYPE_CODEGEN =
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
     8
sig
31246
251a34663242 exported find_shorzes_path
haftmann
parents: 31156
diff changeset
     9
  val find_shortest_path: DatatypeAux.descr -> int -> (string * int) option
31134
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
    10
  val mk_eq_eqns: theory -> string -> (thm * bool) list
30076
f3043dafef5f improved treatment of case certificates
haftmann
parents: 29302
diff changeset
    11
  val mk_case_cert: theory -> string -> thm
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18702
diff changeset
    12
  val setup: theory -> theory
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    13
end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    14
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    15
structure DatatypeCodegen : DATATYPE_CODEGEN =
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    16
struct
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    17
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25505
diff changeset
    18
(** SML code generator **)
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25505
diff changeset
    19
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    20
open Codegen;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    21
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    22
(**** datatype definition ****)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    23
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    24
(* find shortest path to constructor with no recursive arguments *)
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    25
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    26
fun find_nonempty (descr: DatatypeAux.descr) is i =
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    27
  let
31246
251a34663242 exported find_shorzes_path
haftmann
parents: 31156
diff changeset
    28
    val (_, _, constrs) = the (AList.lookup (op =) descr i);
251a34663242 exported find_shorzes_path
haftmann
parents: 31156
diff changeset
    29
    fun arg_nonempty (_, DatatypeAux.DtRec i) = if member (op =) is i
251a34663242 exported find_shorzes_path
haftmann
parents: 31156
diff changeset
    30
          then NONE
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    31
          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
    32
      | arg_nonempty _ = SOME 0;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    33
    fun max xs = Library.foldl
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    34
      (fn (NONE, _) => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    35
        | (SOME i, SOME j) => SOME (Int.max (i, j))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    36
        | (_, NONE) => NONE) (SOME 0, xs);
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    37
    val xs = sort (int_ord o pairself snd)
31246
251a34663242 exported find_shorzes_path
haftmann
parents: 31156
diff changeset
    38
      (map_filter (fn (s, dts) => Option.map (pair s)
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    39
        (max (map (arg_nonempty o DatatypeAux.strip_dtyp) dts))) constrs)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
    40
  in case xs of [] => NONE | x :: _ => SOME x end;
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    41
31246
251a34663242 exported find_shorzes_path
haftmann
parents: 31156
diff changeset
    42
fun find_shortest_path descr i = find_nonempty descr [i] i;
251a34663242 exported find_shorzes_path
haftmann
parents: 31156
diff changeset
    43
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    44
fun add_dt_defs thy defs dep module (descr: DatatypeAux.descr) sorts gr =
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    45
  let
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    46
    val descr' = List.filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr;
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    47
    val rtnames = map (#1 o snd) (List.filter (fn (_, (_, _, cs)) =>
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
    48
      exists (exists DatatypeAux.is_rec_type o snd) cs) descr');
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    49
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    50
    val (_, (tname, _, _)) :: _ = descr';
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    51
    val node_id = tname ^ " (type)";
27398
768da1da59d6 simplified retrieval of theory names of consts and types
haftmann
parents: 27096
diff changeset
    52
    val module' = if_library (thyname_of_type thy tname) module;
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    53
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    54
    fun mk_dtdef prfx [] gr = ([], gr)
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    55
      | mk_dtdef prfx ((_, (tname, dts, cs))::xs) gr =
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    56
          let
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    57
            val tvs = map DatatypeAux.dest_DtTFree dts;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    58
            val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    59
            val ((_, type_id), gr') = mk_type_id module' tname gr;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    60
            val (ps, gr'') = gr' |>
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    61
              fold_map (fn (cname, cargs) =>
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    62
                fold_map (invoke_tycodegen thy defs node_id module' false)
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    63
                  cargs ##>>
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    64
                mk_const_id module' cname) cs';
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    65
            val (rest, gr''') = mk_dtdef "and " xs gr''
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    66
          in
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    67
            (Pretty.block (str prfx ::
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    68
               (if null tvs then [] else
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
    69
                  [mk_tuple (map str tvs), str " "]) @
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
    70
               [str (type_id ^ " ="), Pretty.brk 1] @
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
    71
               List.concat (separate [Pretty.brk 1, str "| "]
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    72
                 (map (fn (ps', (_, cname)) => [Pretty.block
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
    73
                   (str cname ::
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
    74
                    (if null ps' then [] else
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
    75
                     List.concat ([str " of", Pretty.brk 1] ::
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
    76
                       separate [str " *", Pretty.brk 1]
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    77
                         (map single ps'))))]) ps))) :: rest, gr''')
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
    78
          end;
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
    79
25889
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
    80
    fun mk_constr_term cname Ts T ps =
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
    81
      List.concat (separate [str " $", Pretty.brk 1]
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
    82
        ([str ("Const (\"" ^ cname ^ "\","), Pretty.brk 1,
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
    83
          mk_type false (Ts ---> T), str ")"] :: ps));
25889
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
    84
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    85
    fun mk_term_of_def gr prfx [] = []
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    86
      | 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
    87
          let
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
    88
            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
    89
            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
    90
            val T = Type (tname, dts');
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    91
            val rest = mk_term_of_def gr "and " xs;
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30076
diff changeset
    92
            val (_, eqs) = Library.foldl_map (fn (prfx, (cname, Ts)) =>
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
    93
              let val args = map (fn i =>
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
    94
                str ("x" ^ string_of_int i)) (1 upto length Ts)
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
    95
              in ("  | ", Pretty.blk (4,
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
    96
                [str prfx, mk_term_of gr module' false T, Pretty.brk 1,
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    97
                 if null Ts then str (snd (get_const_id gr cname))
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
    98
                 else parens (Pretty.block
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
    99
                   [str (snd (get_const_id gr cname)),
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   100
                    Pretty.brk 1, mk_tuple args]),
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   101
                 str " =", Pretty.brk 1] @
25889
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   102
                 mk_constr_term cname Ts T
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   103
                   (map (fn (x, U) => [Pretty.block [mk_term_of gr module' false U,
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   104
                      Pretty.brk 1, x]]) (args ~~ Ts))))
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   105
              end) (prfx, cs')
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   106
          in eqs @ rest end;
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   107
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   108
    fun mk_gen_of_def gr prfx [] = []
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   109
      | mk_gen_of_def gr prfx ((i, (tname, dts, cs)) :: xs) =
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   110
          let
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   111
            val tvs = map DatatypeAux.dest_DtTFree dts;
25889
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   112
            val Us = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   113
            val T = Type (tname, Us);
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   114
            val (cs1, cs2) =
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   115
              List.partition (exists DatatypeAux.is_rec_type o snd) cs;
31246
251a34663242 exported find_shorzes_path
haftmann
parents: 31156
diff changeset
   116
            val SOME (cname, _) = find_shortest_path descr i;
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   117
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   118
            fun mk_delay p = Pretty.block
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   119
              [str "fn () =>", Pretty.brk 1, p];
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   120
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   121
            fun mk_force p = Pretty.block [p, Pretty.brk 1, str "()"];
25889
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   122
15397
5260ac75e07c Fixed bug in mk_gen_of_def that could cause non-termination of the generator
berghofe
parents: 14981
diff changeset
   123
            fun mk_constr s b (cname, dts) =
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   124
              let
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   125
                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
   126
                    (DatatypeAux.typ_of_dtyp descr sorts dt))
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   127
                  [str (if b andalso DatatypeAux.is_rec_type dt then "0"
15397
5260ac75e07c Fixed bug in mk_gen_of_def that could cause non-termination of the generator
berghofe
parents: 14981
diff changeset
   128
                     else "j")]) dts;
25889
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   129
                val Ts = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   130
                val xs = map str
25889
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   131
                  (DatatypeProp.indexify_names (replicate (length dts) "x"));
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   132
                val ts = map str
25889
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   133
                  (DatatypeProp.indexify_names (replicate (length dts) "t"));
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   134
                val (_, id) = get_const_id gr cname
25889
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   135
              in
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   136
                mk_let
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   137
                  (map2 (fn p => fn q => mk_tuple [p, q]) xs ts ~~ gs)
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   138
                  (mk_tuple
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   139
                    [case xs of
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   140
                       _ :: _ :: _ => Pretty.block
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   141
                         [str id, Pretty.brk 1, mk_tuple xs]
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   142
                     | _ => mk_app false (str id) xs,
25889
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   143
                     mk_delay (Pretty.block (mk_constr_term cname Ts T
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   144
                       (map (single o mk_force) ts)))])
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   145
              end;
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   146
15397
5260ac75e07c Fixed bug in mk_gen_of_def that could cause non-termination of the generator
berghofe
parents: 14981
diff changeset
   147
            fun mk_choice [c] = mk_constr "(i-1)" false c
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   148
              | mk_choice cs = Pretty.block [str "one_of",
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   149
                  Pretty.brk 1, Pretty.blk (1, str "[" ::
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   150
                  List.concat (separate [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
   151
                    (map (single o mk_delay o mk_constr "(i-1)" false) cs)) @
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   152
                  [str "]"]), Pretty.brk 1, str "()"];
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   153
25889
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   154
            val gs = maps (fn s =>
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   155
              let val s' = strip_tname s
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   156
              in [str (s' ^ "G"), str (s' ^ "T")] end) tvs;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   157
            val gen_name = "gen_" ^ snd (get_type_id gr tname)
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   158
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   159
          in
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   160
            Pretty.blk (4, separate (Pretty.brk 1) 
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   161
                (str (prfx ^ gen_name ^
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   162
                   (if null cs1 then "" else "'")) :: gs @
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   163
                 (if null cs1 then [] else [str "i"]) @
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   164
                 [str "j"]) @
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   165
              [str " =", Pretty.brk 1] @
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   166
              (if not (null cs1) andalso not (null cs2)
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   167
               then [str "frequency", Pretty.brk 1,
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   168
                 Pretty.blk (1, [str "[",
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   169
                   mk_tuple [str "i", mk_delay (mk_choice cs1)],
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   170
                   str ",", Pretty.fbrk,
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   171
                   mk_tuple [str "1", mk_delay (mk_choice cs2)],
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   172
                   str "]"]), Pretty.brk 1, str "()"]
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   173
               else if null cs2 then
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   174
                 [Pretty.block [str "(case", Pretty.brk 1,
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   175
                   str "i", Pretty.brk 1, str "of",
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   176
                   Pretty.brk 1, str "0 =>", Pretty.brk 1,
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   177
                   mk_constr "0" true (cname, valOf (AList.lookup (op =) cs cname)),
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   178
                   Pretty.brk 1, str "| _ =>", Pretty.brk 1,
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   179
                   mk_choice cs1, str ")"]]
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   180
               else [mk_choice cs2])) ::
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   181
            (if null cs1 then []
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   182
             else [Pretty.blk (4, separate (Pretty.brk 1) 
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   183
                 (str ("and " ^ gen_name) :: gs @ [str "i"]) @
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   184
               [str " =", Pretty.brk 1] @
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   185
               separate (Pretty.brk 1) (str (gen_name ^ "'") :: gs @
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   186
                 [str "i", str "i"]))]) @
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   187
            mk_gen_of_def gr "and " xs
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   188
          end
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   189
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   190
  in
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   191
    (module', (add_edge_acyclic (node_id, dep) gr
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   192
        handle Graph.CYCLES _ => gr) handle Graph.UNDEF _ =>
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   193
         let
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   194
           val gr1 = add_edge (node_id, dep)
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   195
             (new_node (node_id, (NONE, "", "")) gr);
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   196
           val (dtdef, gr2) = mk_dtdef "datatype " descr' gr1 ;
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   197
         in
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   198
           map_node node_id (K (NONE, module',
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   199
             string_of (Pretty.blk (0, separate Pretty.fbrk dtdef @
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   200
               [str ";"])) ^ "\n\n" ^
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   201
             (if "term_of" mem !mode then
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   202
                string_of (Pretty.blk (0, separate Pretty.fbrk
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   203
                  (mk_term_of_def gr2 "fun " descr') @ [str ";"])) ^ "\n\n"
14104
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   204
              else "") ^
29f726e36e5c Added generator for test case generators.
berghofe
parents: 13754
diff changeset
   205
             (if "test" mem !mode then
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   206
                string_of (Pretty.blk (0, separate Pretty.fbrk
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   207
                  (mk_gen_of_def gr2 "fun " descr') @ [str ";"])) ^ "\n\n"
13754
021cf00435a9 Code generator for datatypes now also generates suitable term_of functions (when
berghofe
parents: 12890
diff changeset
   208
              else ""))) gr2
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   209
         end)
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   210
  end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   211
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   212
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   213
(**** case expressions ****)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   214
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   215
fun pretty_case thy defs dep module brack constrs (c as Const (_, T)) ts gr =
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   216
  let val i = length constrs
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   217
  in if length ts <= i then
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   218
       invoke_codegen thy defs dep module brack (eta_expand c ts (i+1)) gr
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   219
    else
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   220
      let
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   221
        val ts1 = Library.take (i, ts);
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   222
        val t :: ts2 = Library.drop (i, ts);
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30076
diff changeset
   223
        val names = List.foldr OldTerm.add_term_names
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30076
diff changeset
   224
          (map (fst o fst o dest_Var) (List.foldr OldTerm.add_term_vars [] ts1)) ts1;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   225
        val (Ts, dT) = split_last (Library.take (i+1, fst (strip_type T)));
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   226
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   227
        fun pcase [] [] [] gr = ([], gr)
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   228
          | pcase ((cname, cargs)::cs) (t::ts) (U::Us) gr =
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   229
              let
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   230
                val j = length cargs;
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19886
diff changeset
   231
                val xs = Name.variant_list names (replicate j "x");
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   232
                val Us' = Library.take (j, fst (strip_type U));
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   233
                val frees = map Free (xs ~~ Us');
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   234
                val (cp, gr0) = invoke_codegen thy defs dep module false
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   235
                  (list_comb (Const (cname, Us' ---> dT), frees)) gr;
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   236
                val t' = Envir.beta_norm (list_comb (t, frees));
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   237
                val (p, gr1) = invoke_codegen thy defs dep module false t' gr0;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   238
                val (ps, gr2) = pcase cs ts Us gr1;
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   239
              in
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   240
                ([Pretty.block [cp, str " =>", Pretty.brk 1, p]] :: ps, gr2)
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   241
              end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   242
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   243
        val (ps1, gr1) = pcase constrs ts1 Ts gr ;
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   244
        val ps = List.concat (separate [Pretty.brk 1, str "| "] ps1);
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   245
        val (p, gr2) = invoke_codegen thy defs dep module false t gr1;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   246
        val (ps2, gr3) = fold_map (invoke_codegen thy defs dep module true) ts2 gr2;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   247
      in ((if not (null ts2) andalso brack then parens else I)
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   248
        (Pretty.block (separate (Pretty.brk 1)
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   249
          (Pretty.block ([str "(case ", p, str " of",
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   250
             Pretty.brk 1] @ ps @ [str ")"]) :: ps2))), gr3)
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   251
      end
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   252
  end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   253
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   254
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   255
(**** constructors ****)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   256
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   257
fun pretty_constr thy defs dep module brack args (c as Const (s, T)) ts gr =
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   258
  let val i = length args
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   259
  in if i > 1 andalso length ts < i then
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   260
      invoke_codegen thy defs dep module brack (eta_expand c ts i) gr
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   261
     else
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   262
       let
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   263
         val id = mk_qual_id module (get_const_id gr s);
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   264
         val (ps, gr') = fold_map
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   265
           (invoke_codegen thy defs dep module (i = 1)) ts gr;
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   266
       in (case args of
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   267
          _ :: _ :: _ => (if brack then parens else I)
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   268
            (Pretty.block [str id, Pretty.brk 1, mk_tuple ps])
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   269
        | _ => (mk_app brack (str id) ps), gr')
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   270
       end
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   271
  end;
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   272
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   273
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   274
(**** code generators for terms and types ****)
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   275
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   276
fun datatype_codegen thy defs dep module brack t gr = (case strip_comb t of
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   277
   (c as Const (s, T), ts) =>
22778
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   278
     (case DatatypePackage.datatype_of_case thy s of
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   279
        SOME {index, descr, ...} =>
22921
475ff421a6a3 consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents: 22778
diff changeset
   280
          if is_some (get_assoc_code thy (s, T)) then NONE else
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   281
          SOME (pretty_case thy defs dep module brack
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   282
            (#3 (the (AList.lookup op = descr index))) c ts gr )
22778
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   283
      | NONE => case (DatatypePackage.datatype_of_constr thy s, strip_type T) of
28639
9788fb18a142 datatype_codegen now checks name of result type of constructor
berghofe
parents: 28537
diff changeset
   284
        (SOME {index, descr, ...}, (_, U as Type (tyname, _))) =>
22921
475ff421a6a3 consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents: 22778
diff changeset
   285
          if is_some (get_assoc_code thy (s, T)) then NONE else
28639
9788fb18a142 datatype_codegen now checks name of result type of constructor
berghofe
parents: 28537
diff changeset
   286
          let
9788fb18a142 datatype_codegen now checks name of result type of constructor
berghofe
parents: 28537
diff changeset
   287
            val SOME (tyname', _, constrs) = AList.lookup op = descr index;
9788fb18a142 datatype_codegen now checks name of result type of constructor
berghofe
parents: 28537
diff changeset
   288
            val SOME args = AList.lookup op = constrs s
22778
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   289
          in
28639
9788fb18a142 datatype_codegen now checks name of result type of constructor
berghofe
parents: 28537
diff changeset
   290
            if tyname <> tyname' then NONE
9788fb18a142 datatype_codegen now checks name of result type of constructor
berghofe
parents: 28537
diff changeset
   291
            else SOME (pretty_constr thy defs
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   292
              dep module brack args c ts (snd (invoke_tycodegen thy defs dep module false U gr)))
22778
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   293
          end
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   294
      | _ => NONE)
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   295
 | _ => NONE);
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   296
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   297
fun datatype_tycodegen thy defs dep module brack (Type (s, Ts)) gr =
22778
a5b87573f427 Streamlined datatype_codegen function using new datatype_of_case
berghofe
parents: 22746
diff changeset
   298
      (case DatatypePackage.get_datatype thy s of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15397
diff changeset
   299
         NONE => NONE
25889
c93803252748 Test data generation and conversion to terms is now more closely
berghofe
parents: 25864
diff changeset
   300
       | SOME {descr, sorts, ...} =>
22921
475ff421a6a3 consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents: 22778
diff changeset
   301
           if is_some (get_assoc_type thy s) then NONE else
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   302
           let
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   303
             val (ps, gr') = fold_map
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   304
               (invoke_tycodegen thy defs dep module false) Ts gr;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   305
             val (module', gr'') = add_dt_defs thy defs dep module descr sorts gr' ;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   306
             val (tyid, gr''') = mk_type_id module' s gr''
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   307
           in SOME (Pretty.block ((if null Ts then [] else
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26732
diff changeset
   308
               [mk_tuple ps, str " "]) @
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28423
diff changeset
   309
               [str (mk_qual_id module tyid)]), gr''')
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   310
           end)
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 15574
diff changeset
   311
  | datatype_tycodegen _ _ _ _ _ _ _ = NONE;
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   312
18247
b17724cae935 code generator: case expressions, improved name resolving
haftmann
parents: 18220
diff changeset
   313
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25505
diff changeset
   314
(** generic code generator **)
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 20071
diff changeset
   315
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25505
diff changeset
   316
(* case certificates *)
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19008
diff changeset
   317
30076
f3043dafef5f improved treatment of case certificates
haftmann
parents: 29302
diff changeset
   318
fun mk_case_cert thy tyco =
23811
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   319
  let
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   320
    val raw_thms =
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   321
      (#case_rewrites o DatatypePackage.the_datatype thy) tyco;
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   322
    val thms as hd_thm :: _ = raw_thms
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   323
      |> Conjunction.intr_balanced
24976
821628d16552 moved Drule.unvarify to Thm.unvarify (cf. more_thm.ML);
wenzelm
parents: 24845
diff changeset
   324
      |> Thm.unvarify
23811
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   325
      |> Conjunction.elim_balanced (length raw_thms)
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   326
      |> map Simpdata.mk_meta_eq
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   327
      |> map Drule.zero_var_indexes
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   328
    val params = fold_aterms (fn (Free (v, _)) => insert (op =) v
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   329
      | _ => I) (Thm.prop_of hd_thm) [];
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   330
    val rhs = hd_thm
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   331
      |> Thm.prop_of
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   332
      |> Logic.dest_equals
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   333
      |> fst
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   334
      |> Term.strip_comb
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   335
      |> apsnd (fst o split_last)
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   336
      |> list_comb;
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   337
    val lhs = Free (Name.variant params "case", Term.fastype_of rhs);
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   338
    val asm = (Thm.cterm_of thy o Logic.mk_equals) (lhs, rhs);
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   339
  in
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   340
    thms
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   341
    |> Conjunction.intr_balanced
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   342
    |> MetaSimplifier.rewrite_rule [(Thm.symmetric o Thm.assume) asm]
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   343
    |> Thm.implies_intr asm
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   344
    |> Thm.generalize ([], params) 0
25597
34860182b250 moved instance parameter management from class.ML to axclass.ML
haftmann
parents: 25569
diff changeset
   345
    |> AxClass.unoverload thy
23811
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   346
    |> Thm.varifyT
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   347
  end;
b18557301bf9 added function for case certificates
haftmann
parents: 23712
diff changeset
   348
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25505
diff changeset
   349
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25505
diff changeset
   350
(* equality *)
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25505
diff changeset
   351
31134
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   352
fun mk_eq_eqns thy dtco =
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25505
diff changeset
   353
  let
31134
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   354
    val (vs, cos) = DatatypePackage.the_datatype_spec thy dtco;
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   355
    val { descr, index, inject = inject_thms, ... } = DatatypePackage.the_datatype thy dtco;
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   356
    val ty = Type (dtco, map TFree vs);
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   357
    fun mk_eq (t1, t2) = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   358
      $ t1 $ t2;
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   359
    fun true_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.true_const);
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   360
    fun false_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.false_const);
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   361
    val triv_injects = map_filter
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   362
     (fn (c, []) => SOME (HOLogic.mk_Trueprop (true_eq (Const (c, ty), Const (c, ty))))
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   363
       | _ => NONE) cos;
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   364
    fun prep_inject (trueprop $ (equiv $ (_ $ t1 $ t2) $ rhs)) =
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   365
      trueprop $ (equiv $ mk_eq (t1, t2) $ rhs);
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   366
    val injects = map prep_inject (nth (DatatypeProp.make_injs [descr] vs) index);
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   367
    fun prep_distinct (trueprop $ (not $ (_ $ t1 $ t2))) =
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   368
      [trueprop $ false_eq (t1, t2), trueprop $ false_eq (t2, t1)];
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   369
    val distincts = maps prep_distinct (snd (nth (DatatypeProp.make_distincts [descr] vs) index));
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   370
    val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty)));
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   371
    val simpset = Simplifier.context (ProofContext.init thy) (HOL_basic_ss
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   372
      addsimps (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms))
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   373
      addsimprocs [DatatypePackage.distinct_simproc]);
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   374
    fun prove prop = Goal.prove_global thy [] [] prop (K (ALLGOALS (simp_tac simpset)))
31151
1c64b0827ee8 rewrite op = == eq handled by simproc
haftmann
parents: 31134
diff changeset
   375
      |> Simpdata.mk_eq;
31134
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   376
  in map (rpair true o prove) (triv_injects @ injects @ distincts) @ [(prove refl, false)] end;
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   377
31134
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   378
fun add_equality vs dtcos thy =
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   379
  let
28350
715163ec93c0 non left-linear equations for nbe
haftmann
parents: 28084
diff changeset
   380
    fun add_def dtco lthy =
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25889
diff changeset
   381
      let
31134
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   382
        val ty = Type (dtco, map TFree vs);
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25889
diff changeset
   383
        fun mk_side const_name = Const (const_name, ty --> ty --> HOLogic.boolT)
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25889
diff changeset
   384
          $ Free ("x", ty) $ Free ("y", ty);
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25889
diff changeset
   385
        val def = HOLogic.mk_Trueprop (HOLogic.mk_eq
26732
6ea9de67e576 constant HOL.eq now qualified
haftmann
parents: 26513
diff changeset
   386
          (mk_side @{const_name eq_class.eq}, mk_side @{const_name "op ="}));
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25889
diff changeset
   387
        val def' = Syntax.check_term lthy def;
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25889
diff changeset
   388
        val ((_, (_, thm)), lthy') = Specification.definition
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28673
diff changeset
   389
          (NONE, (Attrib.empty_binding, def')) lthy;
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25889
diff changeset
   390
        val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy);
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25889
diff changeset
   391
        val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm;
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25889
diff changeset
   392
      in (thm', lthy') end;
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25889
diff changeset
   393
    fun tac thms = Class.intro_classes_tac []
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25889
diff changeset
   394
      THEN ALLGOALS (ProofContext.fact_tac thms);
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25505
diff changeset
   395
    fun add_eq_thms dtco thy =
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   396
      let
31134
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   397
        val const = AxClass.param_of_inst thy (@{const_name eq_class.eq}, dtco);
24137
8d7896398147 replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
wenzelm
parents: 23811
diff changeset
   398
        val thy_ref = Theory.check_thy thy;
31134
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   399
        fun mk_thms () = rev ((mk_eq_eqns (Theory.deref thy_ref) dtco));
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   400
      in
30076
f3043dafef5f improved treatment of case certificates
haftmann
parents: 29302
diff changeset
   401
        Code.add_eqnl (const, Lazy.lazy mk_thms) thy
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   402
      end;
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   403
  in
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25505
diff changeset
   404
    thy
31134
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   405
    |> TheoryTarget.instantiation (dtcos, vs, [HOLogic.class_eq])
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25889
diff changeset
   406
    |> fold_map add_def dtcos
31134
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   407
    |-> (fn def_thms => Class.prove_instantiation_exit_result (map o Morphism.thm)
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   408
         (fn _ => fn def_thms => tac def_thms) def_thms)
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   409
    |-> (fn def_thms => fold Code.del_eqn def_thms)
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   410
    |> fold add_eq_thms dtcos
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   411
  end;
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   412
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   413
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   414
(* liberal addition of code data for datatypes *)
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   415
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   416
fun mk_constr_consts thy vs dtco cos =
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   417
  let
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   418
    val cs = map (fn (c, tys) => (c, tys ---> Type (dtco, map TFree vs))) cos;
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   419
    val cs' = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) cs;
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31151
diff changeset
   420
  in if is_some (try (Code.constrset_of_consts thy) cs')
31134
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   421
    then SOME cs
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   422
    else NONE
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   423
  end;
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   424
31134
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   425
fun add_all_code dtcos thy =
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   426
  let
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   427
    val (vs :: _, coss) = (split_list o map (DatatypePackage.the_datatype_spec thy)) dtcos;
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   428
    val any_css = map2 (mk_constr_consts thy vs) dtcos coss;
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   429
    val css = if exists is_none any_css then []
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   430
      else map_filter I any_css;
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   431
    val case_rewrites = maps (#case_rewrites o DatatypePackage.the_datatype thy) dtcos;
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   432
    val certs = map (mk_case_cert thy) dtcos;
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   433
  in
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   434
    if null css then thy
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   435
    else thy
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   436
      |> fold Code.add_datatype css
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   437
      |> fold_rev Code.add_default_eqn case_rewrites
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   438
      |> fold Code.add_case certs
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   439
      |> add_equality vs dtcos
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   440
   end;
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   441
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   442
20426
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   443
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   444
(** theory setup **)
9ffea7a8b31c added typecopy_package
haftmann
parents: 20389
diff changeset
   445
19008
14c1b2f5dda4 improved code generator devarification
haftmann
parents: 18963
diff changeset
   446
val setup = 
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   447
  add_codegen "datatype" datatype_codegen
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
   448
  #> add_tycodegen "datatype" datatype_tycodegen
31134
1a5591ecb764 tuned and generalized construction of code equations for eq
haftmann
parents: 30242
diff changeset
   449
  #> DatatypePackage.interpretation add_all_code
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20439
diff changeset
   450
12445
f17eb90bfd16 Code generator for datatypes.
berghofe
parents:
diff changeset
   451
end;