src/HOL/Tools/inductive_codegen.ML
author wenzelm
Wed, 20 Apr 2011 16:49:52 +0200
changeset 42429 7691cc61720a
parent 42428 a2a9018843ae
child 42616 92715b528e78
permissions -rw-r--r--
standardized some ML aliases;
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
42427
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
     9
  val add: string option -> int option -> attribute
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
    10
  val poke_test_fn: (int * int * int -> term list option) -> unit
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
    11
  val test_term: Proof.context -> (term * term list) list -> int list ->
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
    12
    term list option * Quickcheck.report option
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
    13
  val setup: theory -> theory
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
    14
  val quickcheck_setup: theory -> theory
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    15
end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    16
37390
8781d80026fc moved inductive_codegen to place where product type is available; tuned structure name
haftmann
parents: 37198
diff changeset
    17
structure Inductive_Codegen : INDUCTIVE_CODEGEN =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    18
struct
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    19
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    20
(**** theory data ****)
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    21
18930
29d39c10822e replaced Symtab.merge_multi by local merge_rules;
wenzelm
parents: 18928
diff changeset
    22
fun merge_rules tabs =
22642
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
    23
  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
    24
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33338
diff changeset
    25
structure CodegenData = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22791
diff changeset
    26
(
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
    27
  type T =
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    28
    {intros : (thm * (string * int)) list Symtab.table,
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
    29
     graph : unit Graph.T,
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    30
     eqns : (thm * string) list Symtab.table};
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
    31
  val empty =
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
    32
    {intros = Symtab.empty, graph = Graph.empty, eqns = Symtab.empty};
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 15660
diff changeset
    33
  val extend = I;
41472
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 41448
diff changeset
    34
  fun merge
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 41448
diff changeset
    35
    ({intros = intros1, graph = graph1, eqns = eqns1},
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 41448
diff changeset
    36
      {intros = intros2, graph = graph2, eqns = eqns2}) : T =
18930
29d39c10822e replaced Symtab.merge_multi by local merge_rules;
wenzelm
parents: 18928
diff changeset
    37
    {intros = merge_rules (intros1, intros2),
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
    38
     graph = Graph.merge (K true) (graph1, graph2),
18930
29d39c10822e replaced Symtab.merge_multi by local merge_rules;
wenzelm
parents: 18928
diff changeset
    39
     eqns = merge_rules (eqns1, eqns2)};
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22791
diff changeset
    40
);
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    41
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    42
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    43
fun warn thy thm =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    44
  warning ("Inductive_Codegen: Not a proper clause:\n" ^
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    45
    Display.string_of_thm_global thy thm);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
    46
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
    47
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
    48
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    49
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
    50
  let
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    51
    val {intros, graph, eqns} = CodegenData.get thy;
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    52
    fun thyname_of s = (case optmod of
27398
768da1da59d6 simplified retrieval of theory names of consts and types
haftmann
parents: 26975
diff changeset
    53
      NONE => Codegen.thyname_of_const thy s | SOME s => s);
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    54
  in
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    55
    (case Option.map strip_comb (try HOLogic.dest_Trueprop (concl_of thm)) of
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38329
diff changeset
    56
      SOME (Const (@{const_name HOL.eq}, _), [t, _]) =>
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
    57
        (case head_of t of
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
    58
          Const (s, _) =>
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
    59
            CodegenData.put {intros = intros, graph = graph,
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
    60
               eqns = eqns |> Symtab.map_default (s, [])
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
    61
                 (AList.update Thm.eq_thm_prop (thm, thyname_of s))} thy
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    62
        | _ => (warn thy thm; thy))
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    63
    | SOME (Const (s, _), _) =>
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    64
        let
29287
5b0bfd63b5da eliminated OldTerm.(add_)term_consts;
wenzelm
parents: 29270
diff changeset
    65
          val cs = fold Term.add_const_names (Thm.prems_of thm) [];
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    66
          val rules = Symtab.lookup_list intros s;
41489
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41472
diff changeset
    67
          val nparms =
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41472
diff changeset
    68
            (case optnparms of
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41472
diff changeset
    69
              SOME k => k
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41472
diff changeset
    70
            | NONE =>
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41472
diff changeset
    71
                (case rules of
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41472
diff changeset
    72
                  [] =>
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42159
diff changeset
    73
                    (case try (Inductive.the_inductive (Proof_Context.init_global thy)) s of
41489
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41472
diff changeset
    74
                      SOME (_, {raw_induct, ...}) =>
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41472
diff changeset
    75
                        length (Inductive.params_of raw_induct)
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41472
diff changeset
    76
                    | NONE => 0)
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41472
diff changeset
    77
                | xs => snd (snd (List.last xs))))
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    78
        in CodegenData.put
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    79
          {intros = intros |>
22642
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
    80
           Symtab.update (s, (AList.update Thm.eq_thm_prop
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
    81
             (thm, (thyname_of s, nparms)) rules)),
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
    82
           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
    83
           eqns = eqns} thy
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
    84
        end
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    85
    | _ => (warn thy thm; thy))
20897
3f8d2834b2c4 attribute: Context.mapping;
wenzelm
parents: 20071
diff changeset
    86
  end) I);
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    87
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
    88
fun get_clauses thy s =
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    89
  let val {intros, graph, ...} = CodegenData.get thy in
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    90
    (case Symtab.lookup intros s of
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    91
      NONE =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    92
        (case try (Inductive.the_inductive (Proof_Context.init_global thy)) s of
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    93
          NONE => NONE
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    94
        | SOME ({names, ...}, {intrs, raw_induct, ...}) =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    95
            SOME (names, Codegen.thyname_of_const thy s,
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    96
              length (Inductive.params_of raw_induct),
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
    97
              Codegen.preprocess thy intrs))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
    98
    | SOME _ =>
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    99
        let
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
   100
          val SOME names = find_first
22642
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
   101
            (fn xs => member (op =) xs s) (Graph.strong_conn graph);
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
   102
          val intrs as (_, (thyname, nparms)) :: _ =
bfdb29f11eb4 canonical merge operations
haftmann
parents: 22556
diff changeset
   103
            maps (the o Symtab.lookup intros) names;
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   104
        in SOME (names, thyname, nparms, Codegen.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
   105
  end;
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   106
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   107
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   108
(**** check if a term contains only constructor functions ****)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   109
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   110
fun is_constrt thy =
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   111
  let
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31723
diff changeset
   112
    val cnstrs = flat (maps
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   113
      (map (fn (_, (_, _, cs)) => map (apsnd length) cs) o #descr o snd)
33963
977b94b64905 renamed former datatype.ML to datatype_data.ML; datatype.ML provides uniform view on datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33771
diff changeset
   114
      (Symtab.dest (Datatype_Data.get_all thy)));
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   115
    fun check t =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   116
      (case strip_comb t of
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   117
        (Var _, []) => true
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   118
      | (Const (s, _), ts) =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   119
          (case AList.lookup (op =) cnstrs s of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   120
            NONE => false
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   121
          | SOME i => length ts = i andalso forall check ts)
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   122
      | _ => false);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   123
  in check end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   124
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   125
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   126
(**** check if a type is an equality type (i.e. doesn't contain fun) ****)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   127
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   128
fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   129
  | is_eqT _ = true;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   130
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   131
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   132
(**** mode inference ****)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   133
15204
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   134
fun string_of_mode (iss, is) = space_implode " -> " (map
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   135
  (fn NONE => "X"
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   136
    | SOME js => enclose "[" "]" (commas (map string_of_int js)))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   137
       (iss @ [SOME is]));
15204
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   138
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   139
fun print_modes modes = Codegen.message ("Inferred modes:\n" ^
26931
aa226d8405a8 cat_lines;
wenzelm
parents: 26928
diff changeset
   140
  cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   141
    (fn (m, rnd) => string_of_mode m ^
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   142
       (if rnd then " (random)" else "")) ms)) modes));
15204
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   143
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 28537
diff changeset
   144
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
   145
val terms_vs = distinct (op =) o maps term_vs;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   146
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   147
(** collect all Vars in a term (with duplicates!) **)
16861
7446b4be013b tuned fold on terms;
wenzelm
parents: 16645
diff changeset
   148
fun term_vTs tm =
7446b4be013b tuned fold on terms;
wenzelm
parents: 16645
diff changeset
   149
  fold_aterms (fn Var ((x, _), T) => cons (x, T) | _ => I) tm [];
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   150
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   151
fun get_args _ _ [] = ([], [])
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
   152
  | get_args is i (x::xs) = (if member (op =) is i then apfst else apsnd) (cons x)
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   153
      (get_args is (i+1) xs);
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   154
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   155
fun merge xs [] = xs
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   156
  | merge [] ys = ys
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   157
  | 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
   158
      else y::merge (x::xs) ys;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   159
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   160
fun subsets i j = if i <= j then
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   161
       let val is = subsets (i+1) j
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   162
       in merge (map (fn ks => i::ks) is) is end
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   163
     else [[]];
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   164
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   165
fun cprod ([], ys) = []
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   166
  | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys);
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   167
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 29287
diff changeset
   168
fun cprods xss = List.foldr (map op :: o cprod) [[]] xss;
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   169
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   170
datatype mode = Mode of ((int list option list * int list) * bool) * int list * mode option list;
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   171
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   172
fun needs_random (Mode ((_, b), _, ms)) =
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   173
  b orelse exists (fn NONE => false | SOME m => needs_random m) ms;
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   174
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   175
fun modes_of modes t =
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   176
  let
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   177
    val ks = 1 upto length (binder_types (fastype_of t));
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   178
    val default = [Mode ((([], ks), false), ks, [])];
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   179
    fun mk_modes name args = Option.map
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   180
     (maps (fn (m as ((iss, is), _)) =>
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   181
        let
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   182
          val (args1, args2) =
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   183
            if length args < length iss then
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   184
              error ("Too few arguments for inductive predicate " ^ name)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   185
            else chop (length iss) args;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   186
          val k = length args2;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   187
          val prfx = 1 upto k
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   188
        in
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   189
          if not (is_prefix op = prfx is) then [] else
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   190
          let val is' = map (fn i => i - k) (List.drop (is, k))
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   191
          in map (fn x => Mode (m, is', x)) (cprods (map
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   192
            (fn (NONE, _) => [NONE]
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33244
diff changeset
   193
              | (SOME js, arg) => map SOME (filter
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   194
                  (fn Mode (_, js', _) => js=js') (modes_of modes arg)))
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   195
                    (iss ~~ args1)))
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   196
          end
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   197
        end)) (AList.lookup op = modes name)
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   198
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   199
  in
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   200
    (case strip_comb t of
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38329
diff changeset
   201
      (Const (@{const_name HOL.eq}, Type (_, [T, _])), _) =>
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   202
        [Mode ((([], [1]), false), [1], []), Mode ((([], [2]), false), [2], [])] @
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   203
        (if is_eqT T then [Mode ((([], [1, 2]), false), [1, 2], [])] else [])
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   204
    | (Const (name, _), args) => the_default default (mk_modes name args)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   205
    | (Var ((name, _), _), args) => the (mk_modes name args)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   206
    | (Free (name, _), args) => the (mk_modes name args)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   207
    | _ => default)
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   208
  end;
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   209
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 24625
diff changeset
   210
datatype indprem = Prem of term list * term * bool | Sidecond of term;
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   211
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   212
fun missing_vars vs ts = subtract (fn (x, ((y, _), _)) => x = y) vs
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   213
  (fold Term.add_vars ts []);
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   214
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   215
fun monomorphic_vars vs = null (fold (Term.add_tvarsT o snd) vs []);
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   216
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   217
fun mode_ord p = int_ord (pairself (fn (Mode ((_, rnd), _, _), vs) =>
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   218
  length vs + (if null vs then 0 else 1) + (if rnd then 1 else 0)) p);
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   219
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   220
fun select_mode_prem thy modes vs ps =
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   221
  sort (mode_ord o pairself (hd o snd))
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   222
    (filter_out (null o snd) (ps ~~ map
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   223
      (fn Prem (us, t, is_set) => sort mode_ord
42429
7691cc61720a standardized some ML aliases;
wenzelm
parents: 42428
diff changeset
   224
          (map_filter (fn m as Mode (_, is, _) =>
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   225
            let
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   226
              val (in_ts, out_ts) = get_args is 1 us;
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   227
              val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts;
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   228
              val vTs = maps term_vTs out_ts';
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   229
              val dupTs = map snd (duplicates (op =) vTs) @
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   230
                map_filter (AList.lookup (op =) vTs) vs;
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   231
              val missing_vs = missing_vars vs (t :: in_ts @ in_ts')
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   232
            in
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   233
              if forall (is_eqT o fastype_of) in_ts' andalso forall is_eqT dupTs
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   234
                andalso monomorphic_vars missing_vs
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   235
              then SOME (m, missing_vs)
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   236
              else NONE
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   237
            end)
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   238
              (if is_set then [Mode ((([], []), false), [], [])]
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   239
               else modes_of modes t handle Option =>
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   240
                 error ("Bad predicate: " ^ Syntax.string_of_term_global thy t)))
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   241
        | Sidecond t =>
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   242
            let val missing_vs = missing_vars vs [t]
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   243
            in
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   244
              if monomorphic_vars missing_vs
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   245
              then [(Mode ((([], []), false), [], []), missing_vs)]
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   246
              else []
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   247
            end)
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   248
              ps));
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   249
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   250
fun use_random codegen_mode = member (op =) codegen_mode "random_ind";
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   251
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   252
fun check_mode_clause thy codegen_mode arg_vs modes ((iss, is), rnd) (ts, ps) =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   253
  let
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   254
    val modes' = modes @ map_filter
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   255
      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [(([], js), false)]))
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   256
        (arg_vs ~~ iss);
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   257
    fun check_mode_prems vs rnd [] = SOME (vs, rnd)
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   258
      | check_mode_prems vs rnd ps =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   259
          (case select_mode_prem thy modes' vs ps of
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   260
            (x, (m, []) :: _) :: _ =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   261
              check_mode_prems
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   262
                (case x of Prem (us, _, _) => union (op =) vs (terms_vs us) | _ => vs)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   263
                (rnd orelse needs_random m)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   264
                (filter_out (equal x) ps)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   265
          | (_, (_, vs') :: _) :: _ =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   266
              if use_random codegen_mode then
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   267
                check_mode_prems (union (op =) vs (map (fst o fst) vs')) true ps
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   268
              else NONE
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   269
          | _ => NONE);
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   270
    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
   271
    val in_vs = terms_vs in_ts;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   272
  in
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   273
    if forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   274
      forall (is_eqT o fastype_of) in_ts'
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   275
    then
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   276
      (case check_mode_prems (union (op =) arg_vs in_vs) rnd ps of
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   277
        NONE => NONE
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   278
      | SOME (vs, rnd') =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   279
          let val missing_vs = missing_vars vs ts
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   280
          in
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   281
            if null missing_vs orelse
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   282
              use_random codegen_mode andalso monomorphic_vars missing_vs
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   283
            then SOME (rnd' orelse not (null missing_vs))
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   284
            else NONE
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   285
          end)
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   286
    else NONE
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   287
  end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   288
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   289
fun check_modes_pred thy codegen_mode arg_vs preds modes (p, ms) =
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   290
  let val SOME rs = AList.lookup (op =) preds p in
42429
7691cc61720a standardized some ML aliases;
wenzelm
parents: 42428
diff changeset
   291
    (p, map_filter (fn m as (m', _) =>
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   292
      let val xs = map (check_mode_clause thy codegen_mode arg_vs modes m) rs in
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   293
        (case find_index is_none xs of
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   294
          ~1 => SOME (m', exists (fn SOME b => b) xs)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   295
        | i => (Codegen.message ("Clause " ^ string_of_int (i+1) ^ " of " ^
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   296
          p ^ " violates mode " ^ string_of_mode m'); NONE))
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   297
      end) ms)
15204
d15f85347fcb - Inserted additional check for equality types in check_mode_clause that
berghofe
parents: 15126
diff changeset
   298
  end;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   299
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   300
fun fixp f (x : (string * ((int list option list * int list) * bool) list) list) =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   301
  let val y = f x
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   302
  in if x = y then x else fixp f y end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   303
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   304
fun infer_modes thy codegen_mode extra_modes arities arg_vs preds = fixp (fn modes =>
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   305
  map (check_modes_pred thy codegen_mode arg_vs preds (modes @ extra_modes)) modes)
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   306
    (map (fn (s, (ks, k)) => (s, map (rpair false) (cprod (cprods (map
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   307
      (fn NONE => [NONE]
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   308
        | SOME k' => map SOME (subsets 1 k')) ks),
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   309
      subsets 1 k)))) arities);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   310
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   311
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   312
(**** code generation ****)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   313
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   314
fun mk_eq (x::xs) =
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   315
  let
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   316
    fun mk_eqs _ [] = []
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   317
      | mk_eqs a (b :: cs) = Codegen.str (a ^ " = " ^ b) :: mk_eqs b cs;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   318
  in mk_eqs x xs end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   319
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   320
fun mk_tuple xs =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   321
  Pretty.block (Codegen.str "(" ::
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   322
    flat (separate [Codegen.str ",", Pretty.brk 1] (map single xs)) @
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   323
    [Codegen.str ")"]);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   324
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   325
fun mk_v s (names, vs) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   326
  (case AList.lookup (op =) vs s of
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   327
    NONE => (s, (names, (s, [s])::vs))
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   328
  | SOME xs =>
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   329
      let val s' = Name.variant names s
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   330
      in (s', (s'::names, AList.update (op =) (s, s'::xs) vs)) end);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   331
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   332
fun distinct_v (Var ((s, 0), T)) nvs =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   333
      let val (s', nvs') = mk_v s nvs
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   334
      in (Var ((s', 0), T), nvs') end
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   335
  | distinct_v (t $ u) nvs =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   336
      let
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   337
        val (t', nvs') = distinct_v t nvs;
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   338
        val (u', nvs'') = distinct_v u nvs';
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   339
      in (t' $ u', nvs'') end
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   340
  | distinct_v t nvs = (t, nvs);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   341
15061
61a52739cd82 Added simple check that allows code generator to produce code containing
berghofe
parents: 15031
diff changeset
   342
fun is_exhaustive (Var _) = true
37390
8781d80026fc moved inductive_codegen to place where product type is available; tuned structure name
haftmann
parents: 37198
diff changeset
   343
  | is_exhaustive (Const (@{const_name Pair}, _) $ t $ u) =
15061
61a52739cd82 Added simple check that allows code generator to produce code containing
berghofe
parents: 15031
diff changeset
   344
      is_exhaustive t andalso is_exhaustive u
61a52739cd82 Added simple check that allows code generator to produce code containing
berghofe
parents: 15031
diff changeset
   345
  | is_exhaustive _ = false;
61a52739cd82 Added simple check that allows code generator to produce code containing
berghofe
parents: 15031
diff changeset
   346
61a52739cd82 Added simple check that allows code generator to produce code containing
berghofe
parents: 15031
diff changeset
   347
fun compile_match nvs eq_ps out_ps success_p can_fail =
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   348
  let val eqs = flat (separate [Codegen.str " andalso", Pretty.brk 1]
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   349
    (map single (maps (mk_eq o snd) nvs @ eq_ps)));
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   350
  in
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   351
    Pretty.block
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   352
     ([Codegen.str "(fn ", mk_tuple out_ps, Codegen.str " =>", Pretty.brk 1] @
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   353
      (Pretty.block ((if null eqs then [] else Codegen.str "if " ::
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   354
         [Pretty.block eqs, Pretty.brk 1, Codegen.str "then "]) @
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   355
         (success_p ::
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   356
          (if null eqs then [] else [Pretty.brk 1, Codegen.str "else DSeq.empty"]))) ::
15061
61a52739cd82 Added simple check that allows code generator to produce code containing
berghofe
parents: 15031
diff changeset
   357
       (if can_fail then
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   358
          [Pretty.brk 1, Codegen.str "| _ => DSeq.empty)"]
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   359
        else [Codegen.str ")"])))
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   360
  end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   361
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   362
fun modename module s (iss, is) gr =
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38329
diff changeset
   363
  let val (id, gr') = if s = @{const_name HOL.eq} then (("", "equal"), gr)
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   364
    else Codegen.mk_const_id module s gr
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   365
  in (space_implode "__"
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   366
    (Codegen.mk_qual_id module id ::
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   367
      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
   368
  end;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   369
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   370
fun mk_funcomp brack s k p = (if brack then Codegen.parens else I)
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   371
  (Pretty.block [Pretty.block ((if k = 0 then [] else [Codegen.str "("]) @
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   372
    separate (Pretty.brk 1) (Codegen.str s :: replicate k (Codegen.str "|> ???")) @
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   373
    (if k = 0 then [] else [Codegen.str ")"])), Pretty.brk 1, p]);
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   374
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   375
fun compile_expr thy codegen_mode defs dep module brack modes (NONE, t) gr =
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   376
      apfst single (Codegen.invoke_codegen thy codegen_mode defs dep module brack t gr)
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   377
  | compile_expr _ _ _ _ _ _ _ (SOME _, Var ((name, _), _)) gr =
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   378
      ([Codegen.str name], gr)
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   379
  | compile_expr thy codegen_mode
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   380
        defs dep module brack modes (SOME (Mode ((mode, _), _, ms)), t) gr =
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   381
      (case strip_comb t of
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   382
        (Const (name, _), args) =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   383
          if name = @{const_name HOL.eq} orelse AList.defined op = modes name then
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   384
            let
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   385
              val (args1, args2) = chop (length ms) args;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   386
              val ((ps, mode_id), gr') =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   387
                gr |> fold_map
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   388
                  (compile_expr thy codegen_mode defs dep module true modes) (ms ~~ args1)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   389
                ||>> modename module name mode;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   390
               val (ps', gr'') =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   391
                (case mode of
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   392
                   ([], []) => ([Codegen.str "()"], gr')
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   393
                 | _ => fold_map
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   394
                     (Codegen.invoke_codegen thy codegen_mode defs dep module true) args2 gr');
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   395
             in
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   396
              ((if brack andalso not (null ps andalso null ps') then
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   397
                single o Codegen.parens o Pretty.block else I)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   398
                  (flat (separate [Pretty.brk 1]
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   399
                    ([Codegen.str mode_id] :: ps @ map single ps'))), gr')
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   400
             end
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   401
          else
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   402
            apfst (single o mk_funcomp brack "??" (length (binder_types (fastype_of t))))
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   403
              (Codegen.invoke_codegen thy codegen_mode defs dep module true t gr)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   404
      | _ =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   405
        apfst (single o mk_funcomp brack "??" (length (binder_types (fastype_of t))))
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   406
          (Codegen.invoke_codegen thy codegen_mode defs dep module true t gr));
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   407
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   408
fun compile_clause thy codegen_mode 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
   409
  let
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   410
    val modes' = modes @ map_filter
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   411
      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [(([], js), false)]))
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   412
        (arg_vs ~~ iss);
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   413
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   414
    fun check_constrt t (names, eqs) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   415
      if is_constrt thy t then (t, (names, eqs))
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   416
      else
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19861
diff changeset
   417
        let val s = Name.variant names "x";
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   418
        in (Var ((s, 0), fastype_of t), (s::names, (s, t)::eqs)) end;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   419
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   420
    fun compile_eq (s, t) gr =
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   421
      apfst (Pretty.block o cons (Codegen.str (s ^ " = ")) o single)
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   422
        (Codegen.invoke_codegen thy codegen_mode defs dep module false t gr);
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   423
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   424
    val (in_ts, out_ts) = get_args is 1 ts;
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   425
    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
   426
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   427
    fun compile_prems out_ts' vs names [] gr =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   428
          let
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   429
            val (out_ps, gr2) =
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   430
              fold_map (Codegen.invoke_codegen thy codegen_mode defs dep module false)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   431
                out_ts gr;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   432
            val (eq_ps, gr3) = fold_map compile_eq eqs gr2;
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   433
            val (out_ts'', (names', eqs')) = fold_map check_constrt out_ts' (names, []);
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   434
            val (out_ts''', nvs) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   435
              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
   436
            val (out_ps', gr4) =
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   437
              fold_map (Codegen.invoke_codegen thy codegen_mode defs dep module false)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   438
                out_ts''' gr3;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   439
            val (eq_ps', gr5) = fold_map compile_eq eqs' gr4;
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   440
            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts'));
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   441
            val missing_vs = missing_vars vs' out_ts;
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   442
            val final_p = Pretty.block
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   443
              [Codegen.str "DSeq.single", Pretty.brk 1, mk_tuple out_ps]
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   444
          in
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   445
            if null missing_vs then
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   446
              (compile_match (snd nvs) (eq_ps @ eq_ps') out_ps'
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   447
                 final_p (exists (not o is_exhaustive) out_ts'''), gr5)
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   448
            else
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   449
              let
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   450
                val (pat_p, gr6) =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   451
                  Codegen.invoke_codegen thy codegen_mode defs dep module true
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   452
                    (HOLogic.mk_tuple (map Var missing_vs)) gr5;
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   453
                val gen_p =
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   454
                  Codegen.mk_gen gr6 module true [] ""
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   455
                    (HOLogic.mk_tupleT (map snd missing_vs));
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   456
              in
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   457
                (compile_match (snd nvs) eq_ps' out_ps'
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   458
                  (Pretty.block [Codegen.str "DSeq.generator ", gen_p,
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   459
                    Codegen.str " :->", Pretty.brk 1,
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   460
                    compile_match [] eq_ps [pat_p] final_p false])
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   461
                  (exists (not o is_exhaustive) out_ts'''),
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   462
                 gr6)
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   463
              end
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   464
          end
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   465
      | compile_prems out_ts vs names ps gr =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   466
          let
31852
a16bb835e97d explicit Set constructor for code generated for sets
haftmann
parents: 31784
diff changeset
   467
            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   468
            val (out_ts', (names', eqs)) = fold_map check_constrt out_ts (names, []);
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   469
            val (out_ts'', nvs) =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   470
              fold_map distinct_v out_ts' (names', map (fn x => (x, [x])) vs);
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   471
            val (out_ps, gr0) =
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   472
              fold_map (Codegen.invoke_codegen thy codegen_mode defs dep module false)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   473
                out_ts'' gr;
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   474
            val (eq_ps, gr1) = fold_map compile_eq eqs gr0;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   475
          in
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   476
            (case hd (select_mode_prem thy modes' vs' ps) of
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   477
              (p as Prem (us, t, is_set), (mode as Mode (_, js, _), []) :: _) =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   478
                let
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   479
                  val ps' = filter_out (equal p) ps;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   480
                  val (in_ts, out_ts''') = get_args js 1 us;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   481
                  val (in_ps, gr2) =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   482
                    fold_map (Codegen.invoke_codegen thy codegen_mode defs dep module true)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   483
                      in_ts gr1;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   484
                  val (ps, gr3) =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   485
                    if not is_set then
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   486
                      apfst (fn ps => ps @
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   487
                          (if null in_ps then [] else [Pretty.brk 1]) @
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   488
                          separate (Pretty.brk 1) in_ps)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   489
                        (compile_expr thy codegen_mode defs dep module false modes
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   490
                          (SOME mode, t) gr2)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   491
                    else
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   492
                      apfst (fn p =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   493
                        Pretty.breaks [Codegen.str "DSeq.of_list", Codegen.str "(case", p,
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   494
                        Codegen.str "of", Codegen.str "Set", Codegen.str "xs", Codegen.str "=>",
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   495
                        Codegen.str "xs)"])
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   496
                        (*this is a very strong assumption about the generated code!*)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   497
                        (Codegen.invoke_codegen thy codegen_mode defs dep module true t gr2);
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   498
                   val (rest, gr4) = compile_prems out_ts''' vs' (fst nvs) ps' gr3;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   499
                 in
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   500
                   (compile_match (snd nvs) eq_ps out_ps
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   501
                     (Pretty.block (ps @
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   502
                       [Codegen.str " :->", Pretty.brk 1, rest]))
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   503
                       (exists (not o is_exhaustive) out_ts''), gr4)
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   504
                 end
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   505
            | (p as Sidecond t, [(_, [])]) =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   506
                let
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   507
                  val ps' = filter_out (equal p) ps;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   508
                  val (side_p, gr2) =
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   509
                    Codegen.invoke_codegen thy codegen_mode defs dep module true t gr1;
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   510
                  val (rest, gr3) = compile_prems [] vs' (fst nvs) ps' gr2;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   511
                in
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   512
                  (compile_match (snd nvs) eq_ps out_ps
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   513
                    (Pretty.block [Codegen.str "?? ", side_p,
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   514
                      Codegen.str " :->", Pretty.brk 1, rest])
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   515
                    (exists (not o is_exhaustive) out_ts''), gr3)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   516
                end
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   517
            | (_, (_, missing_vs) :: _) =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   518
                let
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   519
                  val T = HOLogic.mk_tupleT (map snd missing_vs);
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   520
                  val (_, gr2) =
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   521
                    Codegen.invoke_tycodegen thy codegen_mode defs dep module false T gr1;
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   522
                  val gen_p = Codegen.mk_gen gr2 module true [] "" T;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   523
                  val (rest, gr3) = compile_prems
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   524
                    [HOLogic.mk_tuple (map Var missing_vs)] vs' (fst nvs) ps gr2;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   525
                in
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   526
                  (compile_match (snd nvs) eq_ps out_ps
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   527
                    (Pretty.block [Codegen.str "DSeq.generator", Pretty.brk 1,
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   528
                      gen_p, Codegen.str " :->", Pretty.brk 1, rest])
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   529
                    (exists (not o is_exhaustive) out_ts''), gr3)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   530
                end)
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   531
          end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   532
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   533
    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
   534
  in
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   535
    (Pretty.block [Codegen.str "DSeq.single", Pretty.brk 1, inp,
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   536
       Codegen.str " :->", Pretty.brk 1, prem_p], gr')
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   537
  end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   538
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   539
fun compile_pred thy codegen_mode 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
   540
  let
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   541
    val xs = map Codegen.str (Name.variant_list arg_vs
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   542
      (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
   543
    val ((cl_ps, mode_id), gr') = gr |>
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   544
      fold_map (fn cl => compile_clause thy codegen_mode defs
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   545
        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
   546
      modename module s mode
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   547
  in
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   548
    (Pretty.block
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   549
      ([Pretty.block (separate (Pretty.brk 1)
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   550
         (Codegen.str (prfx ^ mode_id) ::
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   551
           map Codegen.str arg_vs @
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   552
           (case mode of ([], []) => [Codegen.str "()"] | _ => xs)) @
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   553
         [Codegen.str " ="]),
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   554
        Pretty.brk 1] @
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   555
       flat (separate [Codegen.str " ++", Pretty.brk 1] (map single cl_ps))), (gr', "and "))
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   556
  end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   557
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   558
fun compile_preds thy codegen_mode defs dep module all_vs arg_vs modes preds gr =
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   559
  let val (prs, (gr', _)) = fold_map (fn (s, cls) =>
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   560
    fold_map (fn (mode, _) => fn (gr', prfx') => compile_pred thy codegen_mode defs
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   561
      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
   562
        (((the o AList.lookup (op =) modes) s))) preds (gr, "fun ")
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   563
  in
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   564
    (space_implode "\n\n" (map Codegen.string_of (flat prs)) ^ ";\n\n", gr')
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   565
  end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   566
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   567
(**** processing of introduction rules ****)
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   568
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   569
exception Modes of
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   570
  (string * ((int list option list * int list) * bool) list) list *
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   571
  (string * (int option list * int)) list;
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   572
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   573
fun lookup_modes gr dep = apfst flat (apsnd flat (ListPair.unzip
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   574
  (map ((fn (SOME (Modes x), _, _) => x | _ => ([], [])) o Codegen.get_node gr)
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   575
    (Graph.all_preds (fst gr) [dep]))));
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   576
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   577
fun print_arities arities = Codegen.message ("Arities:\n" ^
26931
aa226d8405a8 cat_lines;
wenzelm
parents: 26928
diff changeset
   578
  cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   579
    space_implode " -> " (map
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   580
      (fn NONE => "X" | SOME k' => string_of_int k')
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   581
        (ks @ [SOME k]))) arities));
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   582
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   583
fun prep_intrs intrs =
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   584
  map (Codegen.rename_term o #prop o rep_thm o Drule.export_without_context) intrs;
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   585
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   586
fun constrain cs [] = []
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   587
  | constrain cs ((s, xs) :: ys) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   588
      (s,
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   589
        (case AList.lookup (op =) cs (s : string) of
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   590
          NONE => xs
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   591
        | SOME xs' => inter (op = o apfst fst) xs' xs)) :: constrain cs ys;
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   592
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   593
fun mk_extra_defs thy codegen_mode defs gr dep names module ts =
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   594
  fold (fn name => fn gr =>
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
   595
    if member (op =) names name then gr
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   596
    else
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   597
      (case get_clauses thy name of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15257
diff changeset
   598
        NONE => gr
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   599
      | SOME (names, thyname, nparms, intrs) =>
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   600
          mk_ind_def thy codegen_mode defs gr dep names
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   601
            (Codegen.if_library codegen_mode thyname module)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   602
            [] (prep_intrs intrs) nparms))
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   603
    (fold Term.add_const_names ts []) gr
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   604
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   605
and mk_ind_def thy codegen_mode defs gr dep names module modecs intrs nparms =
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   606
  Codegen.add_edge_acyclic (hd names, dep) gr handle
24129
591394d9ee66 - Added cycle test to function mk_ind_def to avoid non-termination
berghofe
parents: 23761
diff changeset
   607
    Graph.CYCLES (xs :: _) =>
37390
8781d80026fc moved inductive_codegen to place where product type is available; tuned structure name
haftmann
parents: 37198
diff changeset
   608
      error ("Inductive_Codegen: illegal cyclic dependencies:\n" ^ commas xs)
24129
591394d9ee66 - Added cycle test to function mk_ind_def to avoid non-termination
berghofe
parents: 23761
diff changeset
   609
  | Graph.UNDEF _ =>
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   610
    let
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   611
      val _ $ u = Logic.strip_imp_concl (hd intrs);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   612
      val args = List.take (snd (strip_comb u), nparms);
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32342
diff changeset
   613
      val arg_vs = maps term_vs args;
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   614
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
   615
      fun get_nparms s = if member (op =) names s then SOME nparms else
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   616
        Option.map #3 (get_clauses thy s);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   617
37677
c5a8b612e571 qualified constants Set.member and Set.Collect
haftmann
parents: 37527
diff changeset
   618
      fun dest_prem (_ $ (Const (@{const_name Set.member}, _) $ t $ u)) =
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
   619
            Prem ([t], Envir.beta_eta_contract u, true)
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38329
diff changeset
   620
        | dest_prem (_ $ ((eq as Const (@{const_name HOL.eq}, _)) $ t $ u)) =
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
   621
            Prem ([t, u], eq, false)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   622
        | dest_prem (_ $ t) =
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   623
            (case strip_comb t of
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   624
              (v as Var _, ts) =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   625
                if member (op =) args v then Prem (ts, v, false) else Sidecond t
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
   626
            | (c as Const (s, _), ts) =>
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
   627
                (case get_nparms s of
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
   628
                  NONE => Sidecond t
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
   629
                | SOME k =>
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
   630
                    let val (ts1, ts2) = chop k ts
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
   631
                    in Prem (ts2, list_comb (c, ts1), false) end)
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
   632
            | _ => Sidecond t);
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   633
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   634
      fun add_clause intr (clauses, arities) =
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   635
        let
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   636
          val _ $ t = Logic.strip_imp_concl intr;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   637
          val (Const (name, T), ts) = strip_comb t;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   638
          val (ts1, ts2) = chop nparms ts;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   639
          val prems = map dest_prem (Logic.strip_imp_prems intr);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   640
          val (Ts, Us) = chop nparms (binder_types T)
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   641
        in
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   642
          (AList.update op = (name, these (AList.lookup op = clauses name) @
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   643
             [(ts2, prems)]) clauses,
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   644
           AList.update op = (name, (map (fn U =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   645
              (case strip_type U of
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   646
                (Rs as _ :: _, @{typ bool}) => SOME (length Rs)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   647
              | _ => NONE)) Ts,
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   648
             length Us)) arities)
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   649
        end;
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   650
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   651
      val gr' = mk_extra_defs thy codegen_mode defs
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   652
        (Codegen.add_edge (hd names, dep)
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   653
          (Codegen.new_node (hd names, (NONE, "", "")) gr)) (hd names) names module intrs;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   654
      val (extra_modes, extra_arities) = lookup_modes gr' (hd names);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   655
      val (clauses, arities) = fold add_clause intrs ([], []);
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   656
      val modes = constrain modecs
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   657
        (infer_modes thy codegen_mode extra_modes arities arg_vs clauses);
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   658
      val _ = print_arities arities;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   659
      val _ = print_modes modes;
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   660
      val (s, gr'') =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   661
        compile_preds thy codegen_mode defs (hd names) module (terms_vs intrs)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   662
          arg_vs (modes @ extra_modes) clauses gr';
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   663
    in
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   664
      (Codegen.map_node (hd names)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   665
        (K (SOME (Modes (modes, arities)), module, s)) gr'')
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
   666
    end;
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   667
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   668
fun find_mode gr dep s u modes is =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   669
  (case find_first (fn Mode (_, js, _) => is = js) (modes_of modes u handle Option => []) of
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   670
    NONE =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   671
      Codegen.codegen_error gr dep
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   672
        ("No such mode for " ^ s ^ ": " ^ string_of_mode ([], is))
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   673
  | mode => mode);
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   674
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   675
fun mk_ind_call thy codegen_mode 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
   676
  let
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   677
    val (ts1, ts2) = chop k ts;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   678
    val u = list_comb (Const (s, T), ts1);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   679
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   680
    fun mk_mode (Const (@{const_name dummy_pattern}, _)) ((ts, mode), i) =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   681
          ((ts, mode), i + 1)
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   682
      | mk_mode t ((ts, mode), i) = ((ts @ [t], mode @ [i]), i + 1);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   683
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   684
    val module' = Codegen.if_library codegen_mode thyname module;
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   685
    val gr1 =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   686
      mk_extra_defs thy codegen_mode defs
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   687
        (mk_ind_def thy codegen_mode defs gr dep names module'
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   688
        [] (prep_intrs intrs) k) dep names module' [u];
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   689
    val (modes, _) = lookup_modes gr1 dep;
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   690
    val (ts', is) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   691
      if is_query then fst (fold mk_mode ts2 (([], []), 1))
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   692
      else (ts2, 1 upto length (binder_types T) - k);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   693
    val mode = find_mode gr1 dep s u modes is;
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   694
    val _ = if is_query orelse not (needs_random (the mode)) then ()
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   695
      else warning ("Illegal use of random data generators in " ^ s);
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   696
    val (in_ps, gr2) =
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   697
      fold_map (Codegen.invoke_codegen thy codegen_mode defs dep module true)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   698
        ts' gr1;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   699
    val (ps, gr3) =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   700
      compile_expr thy codegen_mode defs dep module false modes (mode, u) gr2;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   701
  in
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 27809
diff changeset
   702
    (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
   703
       separate (Pretty.brk 1) in_ps), gr3)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   704
  end;
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   705
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   706
fun clause_of_eqn eqn =
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   707
  let
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   708
    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
   709
    val (Const (s, T), ts) = strip_comb t;
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   710
    val (Ts, U) = strip_type T
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   711
  in
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   712
    Codegen.rename_term (Logic.list_implies (prems_of eqn, HOLogic.mk_Trueprop
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   713
      (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
   714
  end;
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   715
42411
ff997038e8eb eliminated Codegen.mode in favour of explicit argument;
wenzelm
parents: 42361
diff changeset
   716
fun mk_fun thy codegen_mode defs name eqns dep module module' gr =
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   717
  (case try (Codegen.get_node gr) name of
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   718
    NONE =>
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   719
      let
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   720
        val clauses = map clause_of_eqn eqns;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   721
        val pname = name ^ "_aux";
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   722
        val arity =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   723
          length (snd (strip_comb (fst (HOLogic.dest_eq
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   724
            (HOLogic.dest_Trueprop (concl_of (hd eqns)))))));
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   725
        val mode = 1 upto arity;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   726
        val ((fun_id, mode_id), gr') = gr |>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   727
          Codegen.mk_const_id module' name ||>>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   728
          modename module' pname ([], mode);
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   729
        val vars = map (fn i => Codegen.str ("x" ^ string_of_int i)) mode;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   730
        val s = Codegen.string_of (Pretty.block
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   731
          [Codegen.mk_app false (Codegen.str ("fun " ^ snd fun_id)) vars, Codegen.str " =",
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   732
           Pretty.brk 1, Codegen.str "DSeq.hd", Pretty.brk 1,
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   733
           Codegen.parens (Pretty.block (separate (Pretty.brk 1) (Codegen.str mode_id ::
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   734
             vars)))]) ^ ";\n\n";
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   735
        val gr'' = mk_ind_def thy codegen_mode defs (Codegen.add_edge (name, dep)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   736
          (Codegen.new_node (name, (NONE, module', s)) gr')) name [pname] module'
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   737
          [(pname, [([], mode)])] clauses 0;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   738
        val (modes, _) = lookup_modes gr'' dep;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   739
        val _ = find_mode gr'' dep pname (head_of (HOLogic.dest_Trueprop
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   740
          (Logic.strip_imp_concl (hd clauses)))) modes mode
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   741
      in (Codegen.mk_qual_id module fun_id, gr'') end
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16861
diff changeset
   742
  | SOME _ =>
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   743
      (Codegen.mk_qual_id module (Codegen.get_const_id gr name),
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   744
        Codegen.add_edge (name, dep) gr));
15031
726dc9146bb1 - Added support for conditional equations whose premises involve
berghofe
parents: 14981
diff changeset
   745
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   746
(* convert n-tuple to nested pairs *)
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   747
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   748
fun conv_ntuple fs ts p =
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   749
  let
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   750
    val k = length fs;
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   751
    val xs = map_range (fn i => Codegen.str ("x" ^ string_of_int i)) (k + 1);
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   752
    val xs' = map (fn Bound i => nth xs (k - i)) ts;
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   753
    fun conv xs js =
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
   754
      if member (op =) fs js then
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   755
        let
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   756
          val (p, xs') = conv xs (1::js);
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   757
          val (q, xs'') = conv xs' (2::js)
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   758
        in (mk_tuple [p, q], xs'') end
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   759
      else (hd xs, tl xs)
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   760
  in
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   761
    if k > 0 then
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   762
      Pretty.block
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   763
        [Codegen.str "DSeq.map (fn", Pretty.brk 1,
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   764
         mk_tuple xs', Codegen.str " =>", Pretty.brk 1, fst (conv xs []),
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   765
         Codegen.str ")", Pretty.brk 1, Codegen.parens p]
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   766
    else p
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   767
  end;
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   768
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   769
fun inductive_codegen thy codegen_mode defs dep module brack t gr =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   770
  (case strip_comb t of
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35021
diff changeset
   771
    (Const (@{const_name Collect}, _), [u]) =>
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   772
      let val (r, Ts, fs) = HOLogic.strip_psplits u in
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   773
        (case strip_comb r of
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   774
          (Const (s, T), ts) =>
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   775
            (case (get_clauses thy s, Codegen.get_assoc_code thy (s, T)) of
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   776
              (SOME (names, thyname, k, intrs), NONE) =>
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   777
                let
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   778
                  val (ts1, ts2) = chop k ts;
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   779
                  val ts2' = map
34962
807f6ce0273d HOLogic.strip_psplits: types are ordered after syntactic appearance, not after corresponding de-Bruin index (closer correspondence to similar strip operations)
haftmann
parents: 33963
diff changeset
   780
                    (fn Bound i => Term.dummy_pattern (nth Ts (length Ts - i - 1)) | t => t) ts2;
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   781
                  val (ots, its) = List.partition is_Bound ts2;
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   782
                  val closed = forall (not o Term.is_open);
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   783
                in
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   784
                  if null (duplicates op = ots) andalso
42083
e1209fc7ecdc added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
wenzelm
parents: 42028
diff changeset
   785
                    closed ts1 andalso closed its
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   786
                  then
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   787
                    let
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   788
                      val (call_p, gr') =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   789
                        mk_ind_call thy codegen_mode defs dep module true
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   790
                          s T (ts1 @ ts2') names thyname k intrs gr;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   791
                    in
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   792
                      SOME ((if brack then Codegen.parens else I) (Pretty.block
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   793
                        [Codegen.str "Set", Pretty.brk 1, Codegen.str "(DSeq.list_of", Pretty.brk 1,
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   794
                         Codegen.str "(", conv_ntuple fs ots call_p, Codegen.str "))"]),
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   795
                         (*this is a very strong assumption about the generated code!*)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   796
                         gr')
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   797
                    end
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   798
                  else NONE
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   799
                end
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   800
            | _ => NONE)
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   801
        | _ => NONE)
23761
9cebbaccf8a2 Improved code generator for Collect.
berghofe
parents: 22921
diff changeset
   802
      end
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   803
  | (Const (s, T), ts) =>
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   804
      (case Symtab.lookup (#eqns (CodegenData.get thy)) s of
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   805
        NONE =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   806
          (case (get_clauses thy s, Codegen.get_assoc_code thy (s, T)) of
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   807
            (SOME (names, thyname, k, intrs), NONE) =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   808
              if length ts < k then NONE else
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   809
                SOME
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   810
                  (let
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   811
                    val (call_p, gr') = mk_ind_call thy codegen_mode defs dep module false
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   812
                      s T (map Term.no_dummy_patterns ts) names thyname k intrs gr
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   813
                   in
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   814
                    (mk_funcomp brack "?!"
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   815
                      (length (binder_types T) - length ts) (Codegen.parens call_p), gr')
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   816
                   end
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   817
                   handle TERM _ =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   818
                    mk_ind_call thy codegen_mode defs dep module true
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   819
                      s T ts names thyname k intrs gr)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   820
          | _ => NONE)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   821
      | SOME eqns =>
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   822
          let
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   823
            val (_, thyname) :: _ = eqns;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   824
            val (id, gr') =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   825
              mk_fun thy codegen_mode defs s (Codegen.preprocess thy (map fst (rev eqns)))
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   826
                dep module (Codegen.if_library codegen_mode thyname module) gr;
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   827
            val (ps, gr'') =
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   828
              fold_map (Codegen.invoke_codegen thy codegen_mode defs dep module true)
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   829
                ts gr';
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   830
          in SOME (Codegen.mk_app brack (Codegen.str id) ps, gr'') end)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21022
diff changeset
   831
  | _ => NONE);
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   832
12557
bb2e4689347e Implemented higher order modes.
berghofe
parents: 12453
diff changeset
   833
val setup =
41448
72ba43b47c7f do not open Codegen;
wenzelm
parents: 40911
diff changeset
   834
  Codegen.add_codegen "inductive" inductive_codegen #>
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   835
  Attrib.setup @{binding code_ind}
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33063
diff changeset
   836
    (Scan.lift (Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) --
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36692
diff changeset
   837
      Scan.option (Args.$$$ "params" |-- Args.colon |-- Parse.nat) >> uncurry add))
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31852
diff changeset
   838
    "introduction rules for executable predicates";
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   839
42427
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   840
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   841
(**** Quickcheck involving inductive predicates ****)
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   842
42427
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   843
structure Result = Proof_Data
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   844
(
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   845
  type T = int * int * int -> term list option;
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   846
  fun init _ = (fn _ => NONE);
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   847
);
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   848
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   849
val get_test_fn = Result.get;
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   850
fun poke_test_fn f = Context.>> (Context.map_proof (Result.put f));
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   851
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   852
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   853
fun strip_imp p =
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   854
  let val (q, r) = HOLogic.dest_imp p
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   855
  in strip_imp r |>> cons q end
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   856
  handle TERM _ => ([], p);
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   857
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   858
fun deepen bound f i =
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   859
  if i > bound then NONE
42428
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   860
  else
a2a9018843ae avoid Display.string_of_thm_without_context;
wenzelm
parents: 42427
diff changeset
   861
    (case f i of
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   862
      NONE => deepen bound f (i + 1)
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   863
    | SOME x => SOME x);
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   864
36001
992839c4be90 static defaults for configuration options;
wenzelm
parents: 35997
diff changeset
   865
val (depth_bound, setup_depth_bound) = Attrib.config_int "ind_quickcheck_depth" (K 10);
992839c4be90 static defaults for configuration options;
wenzelm
parents: 35997
diff changeset
   866
val (depth_start, setup_depth_start) = Attrib.config_int "ind_quickcheck_depth_start" (K 1);
992839c4be90 static defaults for configuration options;
wenzelm
parents: 35997
diff changeset
   867
val (random_values, setup_random_values) = Attrib.config_int "ind_quickcheck_random" (K 5);
992839c4be90 static defaults for configuration options;
wenzelm
parents: 35997
diff changeset
   868
val (size_offset, setup_size_offset) = Attrib.config_int "ind_quickcheck_size_offset" (K 0);
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   869
42159
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42083
diff changeset
   870
fun test_term ctxt [(t, [])] =
42427
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   871
      let
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   872
        val t' = list_abs_free (Term.add_frees t [], t)
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   873
        val thy = Proof_Context.theory_of ctxt;
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   874
        val (xs, p) = strip_abs t';
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   875
        val args' = map_index (fn (i, (_, T)) => ("arg" ^ string_of_int i, T)) xs;
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   876
        val args = map Free args';
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   877
        val (ps, q) = strip_imp p;
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   878
        val Ts = map snd xs;
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   879
        val T = Ts ---> HOLogic.boolT;
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   880
        val rl = Logic.list_implies
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   881
          (map (HOLogic.mk_Trueprop o curry subst_bounds (rev args)) ps @
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   882
           [HOLogic.mk_Trueprop (HOLogic.mk_not (subst_bounds (rev args, q)))],
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   883
           HOLogic.mk_Trueprop (list_comb (Free ("quickcheckp", T), args)));
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   884
        val (_, thy') = Inductive.add_inductive_global
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   885
          {quiet_mode=true, verbose=false, alt_name=Binding.empty, coind=false,
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   886
           no_elim=true, no_ind=false, skip_mono=false, fork_mono=false}
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   887
          [((Binding.name "quickcheckp", T), NoSyn)] []
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   888
          [(Attrib.empty_binding, rl)] [] (Theory.copy thy);
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   889
        val pred = HOLogic.mk_Trueprop (list_comb
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   890
          (Const (Sign.intern_const thy' "quickcheckp", T),
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   891
           map Term.dummy_pattern Ts));
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   892
        val (code, gr) =
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   893
          Codegen.generate_code_i thy' ["term_of", "test", "random_ind"] [] "Generated"
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   894
            [("testf", pred)];
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   895
        val s = "structure Test_Term =\nstruct\n\n" ^
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   896
          cat_lines (map snd code) ^
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   897
          "\nopen Generated;\n\n" ^ Codegen.string_of
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   898
            (Pretty.block [Codegen.str "val () = Inductive_Codegen.poke_test_fn",
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   899
              Pretty.brk 1, Codegen.str "(fn p =>", Pretty.brk 1,
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   900
              Codegen.str "case Seq.pull (testf p) of", Pretty.brk 1,
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   901
              Codegen.str "SOME ",
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   902
              mk_tuple [mk_tuple (map (Codegen.str o fst) args'), Codegen.str "_"],
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   903
              Codegen.str " =>", Pretty.brk 1, Codegen.str "SOME ",
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   904
              Pretty.enum "," "[" "]"
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   905
                (map (fn (s, T) => Pretty.block
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   906
                  [Codegen.mk_term_of gr "Generated" false T, Pretty.brk 1, Codegen.str s]) args'),
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   907
              Pretty.brk 1,
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   908
              Codegen.str "| NONE => NONE);"]) ^
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   909
          "\n\nend;\n";
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   910
        val test_fn =
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   911
          ctxt
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   912
          |> Context.proof_map
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   913
              (ML_Context.exec (fn () => ML_Context.eval_text false Position.none s))
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   914
          |> get_test_fn;
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   915
        val values = Config.get ctxt random_values;
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   916
        val bound = Config.get ctxt depth_bound;
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   917
        val start = Config.get ctxt depth_start;
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   918
        val offset = Config.get ctxt size_offset;
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   919
        fun test [k] = (deepen bound (fn i =>
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   920
          (Output.urgent_message ("Search depth: " ^ string_of_int i);
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   921
           test_fn (i, values, k+offset))) start, NONE);
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   922
      in test end
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   923
  | test_term ctxt [_] = error "Option eval is not supported by tester SML_inductive"
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   924
  | test_term ctxt _ =
5611f178a747 eliminated global references / critical sections via context data;
wenzelm
parents: 42411
diff changeset
   925
      error "Compilation of multiple instances is not supported by tester SML_inductive";
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   926
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   927
val quickcheck_setup =
35997
07bce2802939 use regular Attrib.config instead of low-level Config.declare/Attrib.register_config;
wenzelm
parents: 35382
diff changeset
   928
  setup_depth_bound #>
07bce2802939 use regular Attrib.config instead of low-level Config.declare/Attrib.register_config;
wenzelm
parents: 35382
diff changeset
   929
  setup_depth_start #>
07bce2802939 use regular Attrib.config instead of low-level Config.declare/Attrib.register_config;
wenzelm
parents: 35382
diff changeset
   930
  setup_random_values #>
07bce2802939 use regular Attrib.config instead of low-level Config.declare/Attrib.register_config;
wenzelm
parents: 35382
diff changeset
   931
  setup_size_offset #>
39252
8f176e575a49 changing the container for the quickcheck options to a generic data
bulwahn
parents: 38864
diff changeset
   932
  Context.theory_map (Quickcheck.add_generator ("SML_inductive", test_term));
33771
17926df64f0f Added new counterexample generator SML_inductive for goals involving
berghofe
parents: 33522
diff changeset
   933
11537
e007d35359c3 New code generators for HOL.
berghofe
parents:
diff changeset
   934
end;