src/HOL/Tools/inductive_codegen.ML
author wenzelm
Tue, 17 Nov 2009 14:51:57 +0100
changeset 33726 0878aecbf119
parent 33522 737589bb9bb8
child 33771 17926df64f0f
permissions -rw-r--r--
eliminated slightly odd name space grouping -- now managed by Isar toplevel;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24584
01e83ffa6c54 fixed title
haftmann
parents: 24458
diff changeset
     1
(*  Title:      HOL/Tools/inductive_codegen.ML
11539
0f17da240450 tuned headers;
wenzelm
parents: 11537
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
     3
11539
0f17da240450 tuned headers;
wenzelm
parents: 11537
diff changeset
     4
Code generator for inductive predicates.
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
     5
*)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
     6
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
     7
signature INDUCTIVE_CODEGEN =
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
     8
sig
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
     9
  val add : string option -> int option -> attribute
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18388
diff changeset
    10
  val setup : theory -> theory
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    11
end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    12
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    13
structure InductiveCodegen : INDUCTIVE_CODEGEN =
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    14
struct
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    15
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    16
open Codegen;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    17
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    18
(**** theory data ****)
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    19
18930
29d39c10822e replaced Symtab.merge_multi by local merge_rules;
wenzelm
parents: 18928
diff changeset
    20
fun merge_rules tabs =
22642
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
    21
  Symtab.join (fn _ => AList.merge (Thm.eq_thm_prop) (K true)) tabs;
18930
29d39c10822e replaced Symtab.merge_multi by local merge_rules;
wenzelm
parents: 18928
diff changeset
    22
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33338
diff changeset
    23
structure CodegenData = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22791
diff changeset
    24
(
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
    25
  type T =
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    26
    {intros : (thm * (string * int)) list Symtab.table,
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
    27
     graph : unit Graph.T,
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    28
     eqns : (thm * string) list Symtab.table};
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
    29
  val empty =
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
    30
    {intros = Symtab.empty, graph = Graph.empty, eqns = Symtab.empty};
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 15660
diff changeset
    31
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33338
diff changeset
    32
  fun merge ({intros=intros1, graph=graph1, eqns=eqns1},
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33338
diff changeset
    33
    {intros=intros2, graph=graph2, eqns=eqns2}) : T =
18930
29d39c10822e replaced Symtab.merge_multi by local merge_rules;
wenzelm
parents: 18928
diff changeset
    34
    {intros = merge_rules (intros1, intros2),
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
    35
     graph = Graph.merge (K true) (graph1, graph2),
18930
29d39c10822e replaced Symtab.merge_multi by local merge_rules;
wenzelm
parents: 18928
diff changeset
    36
     eqns = merge_rules (eqns1, eqns2)};
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22791
diff changeset
    37
);
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    38
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    39
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    40
fun warn thm = warning ("InductiveCodegen: Not a proper clause:\n" ^
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31998
diff changeset
    41
  Display.string_of_thm_without_context thm);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    42
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
    43
fun add_node x g = Graph.new_node (x, ()) g handle Graph.DUP _ => g;
14162
f05f299512e9 Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents: 13105
diff changeset
    44
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    45
fun add optmod optnparms = Thm.declaration_attribute (fn thm => Context.mapping (fn thy =>
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    46
  let
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    47
    val {intros, graph, eqns} = CodegenData.get thy;
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    48
    fun thyname_of s = (case optmod of
27398
768da1da59d6 simplified retrieval of theory names of consts and types
haftmann
parents: 26975
diff changeset
    49
      NONE => Codegen.thyname_of_const thy s | SOME s => s);
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    50
  in (case Option.map strip_comb (try HOLogic.dest_Trueprop (concl_of thm)) of
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    51
      SOME (Const ("op =", _), [t, _]) => (case head_of t of
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
    52
        Const (s, _) =>
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    53
          CodegenData.put {intros = intros, graph = graph,
22642
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
    54
             eqns = eqns |> Symtab.map_default (s, [])
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
    55
               (AList.update Thm.eq_thm_prop (thm, thyname_of s))} thy
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    56
      | _ => (warn thm; thy))
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    57
    | SOME (Const (s, _), _) =>
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    58
        let
29287
5b0bfd63b5da eliminated OldTerm.(add_)term_consts;
wenzelm
parents: 29270
diff changeset
    59
          val cs = fold Term.add_const_names (Thm.prems_of thm) [];
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    60
          val rules = Symtab.lookup_list intros s;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    61
          val nparms = (case optnparms of
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    62
            SOME k => k
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    63
          | NONE => (case rules of
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30190
diff changeset
    64
             [] => (case try (Inductive.the_inductive (ProofContext.init thy)) s of
22791
992222f3d2eb Moved function params_of to inductive_package.ML.
berghofe
parents: 22661
diff changeset
    65
                 SOME (_, {raw_induct, ...}) =>
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30190
diff changeset
    66
                   length (Inductive.params_of raw_induct)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    67
               | NONE => 0)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    68
            | xs => snd (snd (snd (split_last xs)))))
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    69
        in CodegenData.put
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    70
          {intros = intros |>
22642
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
    71
           Symtab.update (s, (AList.update Thm.eq_thm_prop
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
    72
             (thm, (thyname_of s, nparms)) rules)),
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
    73
           graph = fold_rev (Graph.add_edge o pair s) cs (fold add_node (s :: cs) graph),
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    74
           eqns = eqns} thy
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    75
        end
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    76
    | _ => (warn thm; thy))
20897
3f8d2834b2c4 attribute: Context.mapping;
wenzelm
parents: 20071
diff changeset
    77
  end) I);
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    78
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    79
fun get_clauses thy s =
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
    80
  let val {intros, graph, ...} = CodegenData.get thy
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17261
diff changeset
    81
  in case Symtab.lookup intros s of
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30190
diff changeset
    82
      NONE => (case try (Inductive.the_inductive (ProofContext.init thy)) s of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
    83
        NONE => NONE
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    84
      | SOME ({names, ...}, {intrs, raw_induct, ...}) =>
27398
768da1da59d6 simplified retrieval of theory names of consts and types
haftmann
parents: 26975
diff changeset
    85
          SOME (names, Codegen.thyname_of_const thy s,
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30190
diff changeset
    86
            length (Inductive.params_of raw_induct),
22661
f3ba63a2663e Removed erroneous application of rev in get_clauses that caused
berghofe
parents: 22642
diff changeset
    87
            preprocess thy intrs))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
    88
    | SOME _ =>
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    89
        let
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    90
          val SOME names = find_first
22642
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
    91
            (fn xs => member (op =) xs s) (Graph.strong_conn graph);
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
    92
          val intrs as (_, (thyname, nparms)) :: _ =
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
    93
            maps (the o Symtab.lookup intros) names;
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
    94
        in SOME (names, thyname, nparms, preprocess thy (map fst (rev intrs))) end
14162
f05f299512e9 Fixed problem with "code ind" attribute that caused code generator to
berghofe
parents: 13105
diff changeset
    95
  end;
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    96
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    97
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    98
(**** check if a term contains only constructor functions ****)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    99
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   100
fun is_constrt thy =
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   101
  let
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31723
diff changeset
   102
    val cnstrs = flat (maps
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   103
      (map (fn (_, (_, _, cs)) => map (apsnd length) cs) o #descr o snd)
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31723
diff changeset
   104
      (Symtab.dest (Datatype.get_all thy)));
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   105
    fun check t = (case strip_comb t of
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   106
        (Var _, []) => true
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   107
      | (Const (s, _), ts) => (case AList.lookup (op =) cnstrs s of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   108
            NONE => false
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   109
          | SOME i => length ts = i andalso forall check ts)
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   110
      | _ => false)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   111
  in check end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   112
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   113
(**** check if a type is an equality type (i.e. doesn't contain fun) ****)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   114
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   115
fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   116
  | is_eqT _ = true;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   117
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   118
(**** mode inference ****)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   119
15204
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   120
fun string_of_mode (iss, is) = space_implode " -> " (map
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   121
  (fn NONE => "X"
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   122
    | SOME js => enclose "[" "]" (commas (map string_of_int js)))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   123
       (iss @ [SOME is]));
15204
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   124
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   125
fun print_modes modes = message ("Inferred modes:\n" ^
26931
aa226d8405a8 cat_lines;
wenzelm
parents: 26928
diff changeset
   126
  cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
15204
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   127
    string_of_mode ms)) modes));
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   128
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 28537
diff changeset
   129
val term_vs = map (fst o fst o dest_Var) o OldTerm.term_vars;
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   130
val terms_vs = distinct (op =) o maps term_vs;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   131
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   132
(** collect all Vars in a term (with duplicates!) **)
16861
7446b4be013b tuned fold on terms;
wenzelm
parents: 16645
diff changeset
   133
fun term_vTs tm =
7446b4be013b tuned fold on terms;
wenzelm
parents: 16645
diff changeset
   134
  fold_aterms (fn Var ((x, _), T) => cons (x, T) | _ => I) tm [];
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   135
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   136
fun get_args _ _ [] = ([], [])
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   137
  | get_args is i (x::xs) = (if i mem is then apfst else apsnd) (cons x)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   138
      (get_args is (i+1) xs);
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   139
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   140
fun merge xs [] = xs
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   141
  | merge [] ys = ys
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   142
  | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   143
      else y::merge (x::xs) ys;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   144
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   145
fun subsets i j = if i <= j then
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   146
       let val is = subsets (i+1) j
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   147
       in merge (map (fn ks => i::ks) is) is end
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   148
     else [[]];
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   149
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   150
fun cprod ([], ys) = []
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   151
  | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys);
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   152
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 29287
diff changeset
   153
fun cprods xss = List.foldr (map op :: o cprod) [[]] xss;
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   154
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   155
datatype mode = Mode of (int list option list * int list) * int list * mode option list;
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   156
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   157
fun modes_of modes t =
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   158
  let
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   159
    val ks = 1 upto length (binder_types (fastype_of t));
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   160
    val default = [Mode (([], ks), ks, [])];
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   161
    fun mk_modes name args = Option.map
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   162
     (maps (fn (m as (iss, is)) =>
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   163
        let
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   164
          val (args1, args2) =
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   165
            if length args < length iss then
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   166
              error ("Too few arguments for inductive predicate " ^ name)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   167
            else chop (length iss) args;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   168
          val k = length args2;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   169
          val prfx = 1 upto k
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   170
        in
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   171
          if not (is_prefix op = prfx is) then [] else
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   172
          let val is' = map (fn i => i - k) (List.drop (is, k))
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   173
          in map (fn x => Mode (m, is', x)) (cprods (map
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   174
            (fn (NONE, _) => [NONE]
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33244
diff changeset
   175
              | (SOME js, arg) => map SOME (filter
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   176
                  (fn Mode (_, js', _) => js=js') (modes_of modes arg)))
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   177
                    (iss ~~ args1)))
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   178
          end
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   179
        end)) (AList.lookup op = modes name)
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   180
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   181
  in (case strip_comb t of
14163
5ffa75ce4919 Improved handling of modes for equality predicate, to avoid ill-typed
berghofe
parents: 14162
diff changeset
   182
      (Const ("op =", Type (_, [T, _])), _) =>
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   183
        [Mode (([], [1]), [1], []), Mode (([], [2]), [2], [])] @
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   184
        (if is_eqT T then [Mode (([], [1, 2]), [1, 2], [])] else [])
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   185
    | (Const (name, _), args) => the_default default (mk_modes name args)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   186
    | (Var ((name, _), _), args) => the (mk_modes name args)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   187
    | (Free (name, _), args) => the (mk_modes name args)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   188
    | _ => default)
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   189
  end;
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   190
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 24625
diff changeset
   191
datatype indprem = Prem of term list * term * bool | Sidecond of term;
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   192
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   193
fun select_mode_prem thy modes vs ps =
19861
620d90091788 tuned Seq/Envir/Unify interfaces;
wenzelm
parents: 19046
diff changeset
   194
  find_first (is_some o snd) (ps ~~ map
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 24625
diff changeset
   195
    (fn Prem (us, t, is_set) => find_first (fn Mode (_, is, _) =>
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   196
          let
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   197
            val (in_ts, out_ts) = get_args is 1 us;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   198
            val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts;
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   199
            val vTs = maps term_vTs out_ts';
18964
67f572e03236 renamed gen_duplicates to duplicates;
wenzelm
parents: 18930
diff changeset
   200
            val dupTs = map snd (duplicates (op =) vTs) @
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   201
              map_filter (AList.lookup (op =) vTs) vs;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   202
          in
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   203
            subset (op =) (terms_vs (in_ts @ in_ts'), vs) andalso
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   204
            forall (is_eqT o fastype_of) in_ts' andalso
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   205
            subset (op =) (term_vs t, vs) andalso
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   206
            forall is_eqT dupTs
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   207
          end)
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 24625
diff changeset
   208
            (if is_set then [Mode (([], []), [], [])]
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 24625
diff changeset
   209
             else modes_of modes t handle Option =>
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26931
diff changeset
   210
               error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   211
      | Sidecond t => if subset (op =) (term_vs t, vs) then SOME (Mode (([], []), [], []))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   212
          else NONE) ps);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   213
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   214
fun check_mode_clause thy arg_vs modes (iss, is) (ts, ps) =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   215
  let
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   216
    val modes' = modes @ map_filter
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   217
      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   218
        (arg_vs ~~ iss);
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   219
    fun check_mode_prems vs [] = SOME vs
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   220
      | check_mode_prems vs ps = (case select_mode_prem thy modes' vs ps of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   221
          NONE => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   222
        | SOME (x, _) => check_mode_prems
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33038
diff changeset
   223
            (case x of Prem (us, _, _) => union (op =) vs (terms_vs us) | _ => vs)
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   224
            (filter_out (equal x) ps));
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   225
    val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (get_args is 1 ts));
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   226
    val in_vs = terms_vs in_ts;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   227
    val concl_vs = terms_vs ts
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   228
  in
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   229
    forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso
15204
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   230
    forall (is_eqT o fastype_of) in_ts' andalso
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33038
diff changeset
   231
    (case check_mode_prems (union (op =) arg_vs in_vs) ps of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   232
       NONE => false
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   233
     | SOME vs => subset (op =) (concl_vs, vs))
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   234
  end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   235
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   236
fun check_modes_pred thy arg_vs preds modes (p, ms) =
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   237
  let val SOME rs = AList.lookup (op =) preds p
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33244
diff changeset
   238
  in (p, filter (fn m => case find_index
15204
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   239
    (not o check_mode_clause thy arg_vs modes m) rs of
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   240
      ~1 => true
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   241
    | i => (message ("Clause " ^ string_of_int (i+1) ^ " of " ^
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   242
      p ^ " violates mode " ^ string_of_mode m); false)) ms)
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   243
  end;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   244
22642
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
   245
fun fixp f (x : (string * (int list option list * int list) list) list) =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   246
  let val y = f x
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   247
  in if x = y then x else fixp f y end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   248
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   249
fun infer_modes thy extra_modes arities arg_vs preds = fixp (fn modes =>
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   250
  map (check_modes_pred thy arg_vs preds (modes @ extra_modes)) modes)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   251
    (map (fn (s, (ks, k)) => (s, cprod (cprods (map
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   252
      (fn NONE => [NONE]
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   253
        | SOME k' => map SOME (subsets 1 k')) ks),
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   254
      subsets 1 k))) arities);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   255
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   256
(**** code generation ****)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   257
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   258
fun mk_eq (x::xs) =
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   259
  let fun mk_eqs _ [] = []
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   260
        | mk_eqs a (b::cs) = str (a ^ " = " ^ b) :: mk_eqs b cs
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   261
  in mk_eqs x xs end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   262
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   263
fun mk_tuple xs = Pretty.block (str "(" ::
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   264
  flat (separate [str ",", Pretty.brk 1] (map single xs)) @
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   265
  [str ")"]);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   266
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   267
fun mk_v s (names, vs) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   268
  (case AList.lookup (op =) vs s of
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   269
    NONE => (s, (names, (s, [s])::vs))
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   270
  | SOME xs =>
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   271
      let val s' = Name.variant names s
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   272
      in (s', (s'::names, AList.update (op =) (s, s'::xs) vs)) end);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   273
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   274
fun distinct_v (Var ((s, 0), T)) nvs =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   275
      let val (s', nvs') = mk_v s nvs
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   276
      in (Var ((s', 0), T), nvs') end
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   277
  | distinct_v (t $ u) nvs =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   278
      let
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   279
        val (t', nvs') = distinct_v t nvs;
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   280
        val (u', nvs'') = distinct_v u nvs';
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   281
      in (t' $ u', nvs'') end
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   282
  | distinct_v t nvs = (t, nvs);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   283
15061
61a52739cd82 Added simple check that allows code generator to produce code containing
berghofe
parents: 15031
diff changeset
   284
fun is_exhaustive (Var _) = true
61a52739cd82 Added simple check that allows code generator to produce code containing
berghofe
parents: 15031
diff changeset
   285
  | is_exhaustive (Const ("Pair", _) $ t $ u) =
61a52739cd82 Added simple check that allows code generator to produce code containing
berghofe
parents: 15031
diff changeset
   286
      is_exhaustive t andalso is_exhaustive u
61a52739cd82 Added simple check that allows code generator to produce code containing
berghofe
parents: 15031
diff changeset
   287
  | is_exhaustive _ = false;
61a52739cd82 Added simple check that allows code generator to produce code containing
berghofe
parents: 15031
diff changeset
   288
61a52739cd82 Added simple check that allows code generator to produce code containing
berghofe
parents: 15031
diff changeset
   289
fun compile_match nvs eq_ps out_ps success_p can_fail =
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   290
  let val eqs = flat (separate [str " andalso", Pretty.brk 1]
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   291
    (map single (maps (mk_eq o snd) nvs @ eq_ps)));
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   292
  in
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   293
    Pretty.block
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   294
     ([str "(fn ", mk_tuple out_ps, str " =>", Pretty.brk 1] @
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   295
      (Pretty.block ((if eqs=[] then [] else str "if " ::
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   296
         [Pretty.block eqs, Pretty.brk 1, str "then "]) @
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   297
         (success_p ::
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   298
          (if eqs=[] then [] else [Pretty.brk 1, str "else DSeq.empty"]))) ::
15061
61a52739cd82 Added simple check that allows code generator to produce code containing
berghofe
parents: 15031
diff changeset
   299
       (if can_fail then
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   300
          [Pretty.brk 1, str "| _ => DSeq.empty)"]
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   301
        else [str ")"])))
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   302
  end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   303
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   304
fun modename module s (iss, is) gr =
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   305
  let val (id, gr') = if s = @{const_name "op ="} then (("", "equal"), gr)
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   306
    else mk_const_id module s gr
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   307
  in (space_implode "__"
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   308
    (mk_qual_id module id ::
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   309
      map (space_implode "_" o map string_of_int) (map_filter I iss @ [is])), gr')
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   310
  end;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   311
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   312
fun mk_funcomp brack s k p = (if brack then parens else I)
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   313
  (Pretty.block [Pretty.block ((if k = 0 then [] else [str "("]) @
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   314
    separate (Pretty.brk 1) (str s :: replicate k (str "|> ???")) @
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   315
    (if k = 0 then [] else [str ")"])), Pretty.brk 1, p]);
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   316
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   317
fun compile_expr thy defs dep module brack modes (NONE, t) gr =
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   318
      apfst single (invoke_codegen thy defs dep module brack t gr)
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   319
  | compile_expr _ _ _ _ _ _ (SOME _, Var ((name, _), _)) gr =
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   320
      ([str name], gr)
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   321
  | compile_expr thy defs dep module brack modes (SOME (Mode (mode, _, ms)), t) gr =
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   322
      (case strip_comb t of
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   323
         (Const (name, _), args) =>
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   324
           if name = @{const_name "op ="} orelse AList.defined op = modes name then
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   325
             let
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   326
               val (args1, args2) = chop (length ms) args;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   327
               val ((ps, mode_id), gr') = gr |> fold_map
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   328
                   (compile_expr thy defs dep module true modes) (ms ~~ args1)
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   329
                   ||>> modename module name mode;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   330
               val (ps', gr'') = (case mode of
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   331
                   ([], []) => ([str "()"], gr')
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   332
                 | _ => fold_map
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   333
                     (invoke_codegen thy defs dep module true) args2 gr')
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   334
             in ((if brack andalso not (null ps andalso null ps') then
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   335
               single o parens o Pretty.block else I)
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   336
                 (flat (separate [Pretty.brk 1]
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   337
                   ([str mode_id] :: ps @ map single ps'))), gr')
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   338
             end
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   339
           else apfst (single o mk_funcomp brack "??" (length (binder_types (fastype_of t))))
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   340
             (invoke_codegen thy defs dep module true t gr)
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   341
       | _ => apfst (single o mk_funcomp brack "??" (length (binder_types (fastype_of t))))
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   342
           (invoke_codegen thy defs dep module true t gr));
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   343
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   344
fun compile_clause thy defs dep module all_vs arg_vs modes (iss, is) (ts, ps) inp gr =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   345
  let
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   346
    val modes' = modes @ map_filter
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   347
      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   348
        (arg_vs ~~ iss);
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   349
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   350
    fun check_constrt t (names, eqs) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   351
      if is_constrt thy t then (t, (names, eqs))
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   352
      else
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19861
diff changeset
   353
        let val s = Name.variant names "x";
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   354
        in (Var ((s, 0), fastype_of t), (s::names, (s, t)::eqs)) end;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   355
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   356
    fun compile_eq (s, t) gr =
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   357
      apfst (Pretty.block o cons (str (s ^ " = ")) o single)
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   358
        (invoke_codegen thy defs dep module false t gr);
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   359
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   360
    val (in_ts, out_ts) = get_args is 1 ts;
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   361
    val (in_ts', (all_vs', eqs)) = fold_map check_constrt in_ts (all_vs, []);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   362
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   363
    fun compile_prems out_ts' vs names [] gr =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   364
          let
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   365
            val (out_ps, gr2) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   366
              fold_map (invoke_codegen thy defs dep module false) out_ts gr;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   367
            val (eq_ps, gr3) = fold_map compile_eq eqs gr2;
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   368
            val (out_ts'', (names', eqs')) = fold_map check_constrt out_ts' (names, []);
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   369
            val (out_ts''', nvs) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   370
              fold_map distinct_v out_ts'' (names', map (fn x => (x, [x])) vs);
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   371
            val (out_ps', gr4) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   372
              fold_map (invoke_codegen thy defs dep module false) out_ts''' gr3;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   373
            val (eq_ps', gr5) = fold_map compile_eq eqs' gr4;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   374
          in
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   375
            (compile_match (snd nvs) (eq_ps @ eq_ps') out_ps'
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   376
              (Pretty.block [str "DSeq.single", Pretty.brk 1, mk_tuple out_ps])
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   377
              (exists (not o is_exhaustive) out_ts'''), gr5)
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   378
          end
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   379
      | compile_prems out_ts vs names ps gr =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   380
          let
31852
a16bb835e97d explicit Set constructor for code generated for sets
haftmann
parents: 31784
diff changeset
   381
            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   382
            val SOME (p, mode as SOME (Mode (_, js, _))) = select_mode_prem thy modes' vs' ps;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   383
            val ps' = filter_out (equal p) ps;
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   384
            val (out_ts', (names', eqs)) = fold_map check_constrt out_ts (names, []);
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   385
            val (out_ts'', nvs) = fold_map distinct_v out_ts' (names', map (fn x => (x, [x])) vs);
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   386
            val (out_ps, gr0) = fold_map (invoke_codegen thy defs dep module false) out_ts'' gr;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   387
            val (eq_ps, gr1) = fold_map compile_eq eqs gr0;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   388
          in
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   389
            (case p of
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 24625
diff changeset
   390
               Prem (us, t, is_set) =>
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   391
                 let
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   392
                   val (in_ts, out_ts''') = get_args js 1 us;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   393
                   val (in_ps, gr2) = fold_map
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   394
                     (invoke_codegen thy defs dep module true) in_ts gr1;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   395
                   val (ps, gr3) =
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 24625
diff changeset
   396
                     if not is_set then
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   397
                       apfst (fn ps => ps @
24129
591394d9ee66 - Added cycle test to function mk_ind_def to avoid non-termination
berghofe
parents: 23761
diff changeset
   398
                           (if null in_ps then [] else [Pretty.brk 1]) @
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   399
                           separate (Pretty.brk 1) in_ps)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   400
                         (compile_expr thy defs dep module false modes
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   401
                           (mode, t) gr2)
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 24625
diff changeset
   402
                     else
31852
a16bb835e97d explicit Set constructor for code generated for sets
haftmann
parents: 31784
diff changeset
   403
                       apfst (fn p => Pretty.breaks [str "DSeq.of_list", str "(case", p,
a16bb835e97d explicit Set constructor for code generated for sets
haftmann
parents: 31784
diff changeset
   404
                         str "of", str "Set", str "xs", str "=>", str "xs)"])
a16bb835e97d explicit Set constructor for code generated for sets
haftmann
parents: 31784
diff changeset
   405
                         (*this is a very strong assumption about the generated code!*)
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   406
                           (invoke_codegen thy defs dep module true t gr2);
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   407
                   val (rest, gr4) = compile_prems out_ts''' vs' (fst nvs) ps' gr3;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   408
                 in
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   409
                   (compile_match (snd nvs) eq_ps out_ps
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   410
                      (Pretty.block (ps @
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   411
                         [str " :->", Pretty.brk 1, rest]))
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   412
                      (exists (not o is_exhaustive) out_ts''), gr4)
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   413
                 end
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   414
             | Sidecond t =>
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   415
                 let
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   416
                   val (side_p, gr2) = invoke_codegen thy defs dep module true t gr1;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   417
                   val (rest, gr3) = compile_prems [] vs' (fst nvs) ps' gr2;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   418
                 in
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   419
                   (compile_match (snd nvs) eq_ps out_ps
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   420
                      (Pretty.block [str "?? ", side_p,
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   421
                        str " :->", Pretty.brk 1, rest])
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   422
                      (exists (not o is_exhaustive) out_ts''), gr3)
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   423
                 end)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   424
          end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   425
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   426
    val (prem_p, gr') = compile_prems in_ts' arg_vs all_vs' ps gr ;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   427
  in
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   428
    (Pretty.block [str "DSeq.single", Pretty.brk 1, inp,
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   429
       str " :->", Pretty.brk 1, prem_p], gr')
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   430
  end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   431
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   432
fun compile_pred thy defs dep module prfx all_vs arg_vs modes s cls mode gr =
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   433
  let
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   434
    val xs = map str (Name.variant_list arg_vs
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   435
      (map (fn i => "x" ^ string_of_int i) (snd mode)));
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   436
    val ((cl_ps, mode_id), gr') = gr |>
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   437
      fold_map (fn cl => compile_clause thy defs
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   438
        dep module all_vs arg_vs modes mode cl (mk_tuple xs)) cls ||>>
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   439
      modename module s mode
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   440
  in
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   441
    (Pretty.block
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   442
      ([Pretty.block (separate (Pretty.brk 1)
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   443
         (str (prfx ^ mode_id) ::
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   444
           map str arg_vs @
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   445
           (case mode of ([], []) => [str "()"] | _ => xs)) @
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   446
         [str " ="]),
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   447
        Pretty.brk 1] @
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   448
       flat (separate [str " ++", Pretty.brk 1] (map single cl_ps))), (gr', "and "))
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   449
  end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   450
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   451
fun compile_preds thy defs dep module all_vs arg_vs modes preds gr =
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   452
  let val (prs, (gr', _)) = fold_map (fn (s, cls) =>
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   453
    fold_map (fn mode => fn (gr', prfx') => compile_pred thy defs
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   454
      dep module prfx' all_vs arg_vs modes s cls mode gr')
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   455
        (((the o AList.lookup (op =) modes) s))) preds (gr, "fun ")
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   456
  in
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   457
    (space_implode "\n\n" (map string_of (flat prs)) ^ ";\n\n", gr')
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   458
  end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   459
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   460
(**** processing of introduction rules ****)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   461
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   462
exception Modes of
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   463
  (string * (int list option list * int list) list) list *
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   464
  (string * (int option list * int)) list;
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   465
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   466
fun lookup_modes gr dep = apfst flat (apsnd flat (ListPair.unzip
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   467
  (map ((fn (SOME (Modes x), _, _) => x | _ => ([], [])) o get_node gr)
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   468
    (Graph.all_preds (fst gr) [dep]))));
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   469
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   470
fun print_arities arities = message ("Arities:\n" ^
26931
aa226d8405a8 cat_lines;
wenzelm
parents: 26928
diff changeset
   471
  cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   472
    space_implode " -> " (map
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   473
      (fn NONE => "X" | SOME k' => string_of_int k')
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   474
        (ks @ [SOME k]))) arities));
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   475
32957
675c0c7e6a37 explicitly qualify Drule.standard;
wenzelm
parents: 32952
diff changeset
   476
fun prep_intrs intrs = map (rename_term o #prop o rep_thm o Drule.standard) intrs;
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   477
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   478
fun constrain cs [] = []
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   479
  | constrain cs ((s, xs) :: ys) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   480
      (s,
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   481
        case AList.lookup (op =) cs (s : string) of
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   482
          NONE => xs
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   483
        | SOME xs' => inter (op =) xs' xs) :: constrain cs ys;
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   484
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   485
fun mk_extra_defs thy defs gr dep names module ts =
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   486
  fold (fn name => fn gr =>
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   487
    if name mem names then gr
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   488
    else
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   489
      (case get_clauses thy name of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   490
        NONE => gr
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   491
      | SOME (names, thyname, nparms, intrs) =>
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   492
          mk_ind_def thy defs gr dep names (if_library thyname module)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   493
            [] (prep_intrs intrs) nparms))
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   494
    (fold Term.add_const_names ts []) gr
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   495
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   496
and mk_ind_def thy defs gr dep names module modecs intrs nparms =
24129
591394d9ee66 - Added cycle test to function mk_ind_def to avoid non-termination
berghofe
parents: 23761
diff changeset
   497
  add_edge_acyclic (hd names, dep) gr handle
591394d9ee66 - Added cycle test to function mk_ind_def to avoid non-termination
berghofe
parents: 23761
diff changeset
   498
    Graph.CYCLES (xs :: _) =>
591394d9ee66 - Added cycle test to function mk_ind_def to avoid non-termination
berghofe
parents: 23761
diff changeset
   499
      error ("InductiveCodegen: illegal cyclic dependencies:\n" ^ commas xs)
591394d9ee66 - Added cycle test to function mk_ind_def to avoid non-termination
berghofe
parents: 23761
diff changeset
   500
  | Graph.UNDEF _ =>
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   501
    let
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   502
      val _ $ u = Logic.strip_imp_concl (hd intrs);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   503
      val args = List.take (snd (strip_comb u), nparms);
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   504
      val arg_vs = maps term_vs args;
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   505
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   506
      fun get_nparms s = if s mem names then SOME nparms else
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   507
        Option.map #3 (get_clauses thy s);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   508
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 24625
diff changeset
   509
      fun dest_prem (_ $ (Const ("op :", _) $ t $ u)) = Prem ([t], Envir.beta_eta_contract u, true)
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 24625
diff changeset
   510
        | dest_prem (_ $ ((eq as Const ("op =", _)) $ t $ u)) = Prem ([t, u], eq, false)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   511
        | dest_prem (_ $ t) =
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   512
            (case strip_comb t of
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 24625
diff changeset
   513
               (v as Var _, ts) => if v mem args then Prem (ts, v, false) else Sidecond t
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   514
             | (c as Const (s, _), ts) => (case get_nparms s of
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   515
                 NONE => Sidecond t
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   516
               | SOME k =>
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   517
                   let val (ts1, ts2) = chop k ts
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 24625
diff changeset
   518
                   in Prem (ts2, list_comb (c, ts1), false) end)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   519
             | _ => Sidecond t);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   520
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   521
      fun add_clause intr (clauses, arities) =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   522
        let
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   523
          val _ $ t = Logic.strip_imp_concl intr;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   524
          val (Const (name, T), ts) = strip_comb t;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   525
          val (ts1, ts2) = chop nparms ts;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   526
          val prems = map dest_prem (Logic.strip_imp_prems intr);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   527
          val (Ts, Us) = chop nparms (binder_types T)
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   528
        in
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   529
          (AList.update op = (name, these (AList.lookup op = clauses name) @
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   530
             [(ts2, prems)]) clauses,
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   531
           AList.update op = (name, (map (fn U => (case strip_type U of
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   532
                 (Rs as _ :: _, Type ("bool", [])) => SOME (length Rs)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   533
               | _ => NONE)) Ts,
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   534
             length Us)) arities)
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   535
        end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   536
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
   537
      val gr' = mk_extra_defs thy defs
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   538
        (add_edge (hd names, dep)
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   539
          (new_node (hd names, (NONE, "", "")) gr)) (hd names) names module intrs;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   540
      val (extra_modes, extra_arities) = lookup_modes gr' (hd names);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   541
      val (clauses, arities) = fold add_clause intrs ([], []);
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   542
      val modes = constrain modecs
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   543
        (infer_modes thy extra_modes arities arg_vs clauses);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   544
      val _ = print_arities arities;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   545
      val _ = print_modes modes;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   546
      val (s, gr'') = compile_preds thy defs (hd names) module (terms_vs intrs)
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   547
        arg_vs (modes @ extra_modes) clauses gr';
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   548
    in
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   549
      (map_node (hd names)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   550
        (K (SOME (Modes (modes, arities)), module, s)) gr'')
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
   551
    end;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   552
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   553
fun find_mode gr dep s u modes is = (case find_first (fn Mode (_, js, _) => is=js)
15660
255055554c67 handle Option instead of OPTION;
wenzelm
parents: 15574
diff changeset
   554
  (modes_of modes u handle Option => []) of
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   555
     NONE => codegen_error gr dep
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   556
       ("No such mode for " ^ s ^ ": " ^ string_of_mode ([], is))
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   557
   | mode => mode);
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   558
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   559
fun mk_ind_call thy defs dep module is_query s T ts names thyname k intrs gr =
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   560
  let
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   561
    val (ts1, ts2) = chop k ts;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   562
    val u = list_comb (Const (s, T), ts1);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   563
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   564
    fun mk_mode (Const ("dummy_pattern", _)) ((ts, mode), i) = ((ts, mode), i + 1)
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   565
      | mk_mode t ((ts, mode), i) = ((ts @ [t], mode @ [i]), i + 1);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   566
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   567
    val module' = if_library thyname module;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   568
    val gr1 = mk_extra_defs thy defs
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   569
      (mk_ind_def thy defs gr dep names module'
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   570
      [] (prep_intrs intrs) k) dep names module' [u];
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   571
    val (modes, _) = lookup_modes gr1 dep;
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   572
    val (ts', is) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   573
      if is_query then fst (fold mk_mode ts2 (([], []), 1))
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   574
      else (ts2, 1 upto length (binder_types T) - k);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   575
    val mode = find_mode gr1 dep s u modes is;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   576
    val (in_ps, gr2) = fold_map (invoke_codegen thy defs dep module true) ts' gr1;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   577
    val (ps, gr3) = compile_expr thy defs dep module false modes (mode, u) gr2;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   578
  in
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   579
    (Pretty.block (ps @ (if null in_ps then [] else [Pretty.brk 1]) @
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   580
       separate (Pretty.brk 1) in_ps), gr3)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   581
  end;
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   582
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   583
fun clause_of_eqn eqn =
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   584
  let
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   585
    val (t, u) = HOLogic.dest_eq (HOLogic.dest_Trueprop (concl_of eqn));
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   586
    val (Const (s, T), ts) = strip_comb t;
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   587
    val (Ts, U) = strip_type T
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   588
  in
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   589
    rename_term (Logic.list_implies (prems_of eqn, HOLogic.mk_Trueprop
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   590
      (list_comb (Const (s ^ "_aux", Ts @ [U] ---> HOLogic.boolT), ts @ [u]))))
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   591
  end;
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   592
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   593
fun mk_fun thy defs name eqns dep module module' gr =
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   594
  case try (get_node gr) name of
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   595
    NONE =>
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   596
    let
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   597
      val clauses = map clause_of_eqn eqns;
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   598
      val pname = name ^ "_aux";
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   599
      val arity = length (snd (strip_comb (fst (HOLogic.dest_eq
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   600
        (HOLogic.dest_Trueprop (concl_of (hd eqns)))))));
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   601
      val mode = 1 upto arity;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   602
      val ((fun_id, mode_id), gr') = gr |>
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   603
        mk_const_id module' name ||>>
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   604
        modename module' pname ([], mode);
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   605
      val vars = map (fn i => str ("x" ^ string_of_int i)) mode;
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   606
      val s = string_of (Pretty.block
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   607
        [mk_app false (str ("fun " ^ snd fun_id)) vars, str " =",
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   608
         Pretty.brk 1, str "DSeq.hd", Pretty.brk 1,
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   609
         parens (Pretty.block (separate (Pretty.brk 1) (str mode_id ::
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   610
           vars)))]) ^ ";\n\n";
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   611
      val gr'' = mk_ind_def thy defs (add_edge (name, dep)
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   612
        (new_node (name, (NONE, module', s)) gr')) name [pname] module'
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   613
        [(pname, [([], mode)])] clauses 0;
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   614
      val (modes, _) = lookup_modes gr'' dep;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   615
      val _ = find_mode gr'' dep pname (head_of (HOLogic.dest_Trueprop
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   616
        (Logic.strip_imp_concl (hd clauses)))) modes mode
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   617
    in (mk_qual_id module fun_id, gr'') end
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   618
  | SOME _ =>
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   619
      (mk_qual_id module (get_const_id gr name), add_edge (name, dep) gr);
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   620
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   621
(* convert n-tuple to nested pairs *)
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   622
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   623
fun conv_ntuple fs ts p =
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   624
  let
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   625
    val k = length fs;
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33050
diff changeset
   626
    val xs = map_range (fn i => str ("x" ^ string_of_int i)) (k + 1);
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   627
    val xs' = map (fn Bound i => nth xs (k - i)) ts;
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   628
    fun conv xs js =
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   629
      if js mem fs then
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   630
        let
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   631
          val (p, xs') = conv xs (1::js);
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   632
          val (q, xs'') = conv xs' (2::js)
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   633
        in (mk_tuple [p, q], xs'') end
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   634
      else (hd xs, tl xs)
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   635
  in
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   636
    if k > 0 then
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   637
      Pretty.block
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   638
        [str "DSeq.map (fn", Pretty.brk 1,
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   639
         mk_tuple xs', str " =>", Pretty.brk 1, fst (conv xs []),
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26939
diff changeset
   640
         str ")", Pretty.brk 1, parens p]
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   641
    else p
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   642
  end;
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   643
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   644
fun inductive_codegen thy defs dep module brack t gr  = (case strip_comb t of
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   645
    (Const ("Collect", _), [u]) =>
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
   646
      let val (r, Ts, fs) = HOLogic.strip_psplits u
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   647
      in case strip_comb r of
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   648
          (Const (s, T), ts) =>
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   649
            (case (get_clauses thy s, get_assoc_code thy (s, T)) of
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   650
              (SOME (names, thyname, k, intrs), NONE) =>
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   651
                let
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   652
                  val (ts1, ts2) = chop k ts;
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   653
                  val ts2' = map
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   654
                    (fn Bound i => Term.dummy_pattern (nth Ts i) | t => t) ts2;
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   655
                  val (ots, its) = List.partition is_Bound ts2;
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   656
                  val no_loose = forall (fn t => not (loose_bvar (t, 0)))
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   657
                in
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   658
                  if null (duplicates op = ots) andalso
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   659
                    no_loose ts1 andalso no_loose its
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   660
                  then
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   661
                    let val (call_p, gr') = mk_ind_call thy defs dep module true
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   662
                      s T (ts1 @ ts2') names thyname k intrs gr 
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   663
                    in SOME ((if brack then parens else I) (Pretty.block
31852
a16bb835e97d explicit Set constructor for code generated for sets
haftmann
parents: 31784
diff changeset
   664
                      [str "Set", Pretty.brk 1, str "(DSeq.list_of", Pretty.brk 1, str "(",
a16bb835e97d explicit Set constructor for code generated for sets
haftmann
parents: 31784
diff changeset
   665
                       conv_ntuple fs ots call_p, str "))"]),
a16bb835e97d explicit Set constructor for code generated for sets
haftmann
parents: 31784
diff changeset
   666
                       (*this is a very strong assumption about the generated code!*)
a16bb835e97d explicit Set constructor for code generated for sets
haftmann
parents: 31784
diff changeset
   667
                       gr')
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   668
                    end
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   669
                  else NONE
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   670
                end
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   671
            | _ => NONE)
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   672
        | _ => NONE
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   673
      end
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   674
  | (Const (s, T), ts) => (case Symtab.lookup (#eqns (CodegenData.get thy)) s of
22921
475ff421a6a3 consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents: 22846
diff changeset
   675
      NONE => (case (get_clauses thy s, get_assoc_code thy (s, T)) of
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   676
        (SOME (names, thyname, k, intrs), NONE) =>
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   677
          if length ts < k then NONE else SOME
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   678
            (let val (call_p, gr') = mk_ind_call thy defs dep module false
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   679
               s T (map Term.no_dummy_patterns ts) names thyname k intrs gr
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   680
             in (mk_funcomp brack "?!"
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   681
               (length (binder_types T) - length ts) (parens call_p), gr')
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   682
             end handle TERM _ => mk_ind_call thy defs dep module true
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   683
               s T ts names thyname k intrs gr )
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   684
      | _ => NONE)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   685
    | SOME eqns =>
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   686
        let
22642
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
   687
          val (_, thyname) :: _ = eqns;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   688
          val (id, gr') = mk_fun thy defs s (preprocess thy (map fst (rev eqns)))
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   689
            dep module (if_library thyname module) gr;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   690
          val (ps, gr'') = fold_map
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   691
            (invoke_codegen thy defs dep module true) ts gr';
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   692
        in SOME (mk_app brack (str id) ps, gr'')
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   693
        end)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   694
  | _ => NONE);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   695
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   696
val setup =
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18388
diff changeset
   697
  add_codegen "inductive" inductive_codegen #>
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   698
  Attrib.setup @{binding code_ind}
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   699
    (Scan.lift (Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) --
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   700
      Scan.option (Args.$$$ "params" |-- Args.colon |-- OuterParse.nat) >> uncurry add))
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31852
diff changeset
   701
    "introduction rules for executable predicates";
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   702
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   703
end;