src/HOL/Tools/recfun_codegen.ML
author haftmann
Mon, 10 Aug 2009 12:24:49 +0200
changeset 32353 0ac26087464b
parent 32345 4da4fa060bb6
child 32358 98c00ee9e786
permissions -rw-r--r--
moved all technical processing of code equations to code_thingol.ML
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24584
01e83ffa6c54 fixed title
haftmann
parents: 24219
diff changeset
     1
(*  Title:      HOL/Tools/recfun_codegen.ML
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
     3
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
     4
Code generator for recursive functions.
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
     5
*)
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
     6
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
     7
signature RECFUN_CODEGEN =
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
     8
sig
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18702
diff changeset
     9
  val setup: theory -> theory
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    10
end;
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    11
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    12
structure RecfunCodegen : RECFUN_CODEGEN =
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    13
struct
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    14
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    15
open Codegen;
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    16
28522
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 28370
diff changeset
    17
structure ModuleData = TheoryDataFun
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    18
(
28522
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 28370
diff changeset
    19
  type T = string Symtab.table;
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    20
  val empty = Symtab.empty;
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    21
  val copy = I;
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 15700
diff changeset
    22
  val extend = I;
28522
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 28370
diff changeset
    23
  fun merge _ = Symtab.merge (K true);
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    24
);
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    25
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31962
diff changeset
    26
fun add_thm_target module_name thm thy =
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31962
diff changeset
    27
  let
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31962
diff changeset
    28
    val (thm', _) = Code.mk_eqn thy (thm, true)
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31962
diff changeset
    29
  in
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31962
diff changeset
    30
    thy
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31962
diff changeset
    31
    |> ModuleData.map (Symtab.update (fst (Code.const_typ_eqn thy thm'), module_name))
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31962
diff changeset
    32
  end;
24624
b8383b1bbae3 distinction between regular and default code theorems
haftmann
parents: 24584
diff changeset
    33
28522
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 28370
diff changeset
    34
fun retrieve_equations thy (c, T) = if c = @{const_name "op ="} then NONE else
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 28370
diff changeset
    35
  let
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 28370
diff changeset
    36
    val c' = AxClass.unoverload_const thy (c, T);
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 28370
diff changeset
    37
    val opt_name = Symtab.lookup (ModuleData.get thy) c';
32353
0ac26087464b moved all technical processing of code equations to code_thingol.ML
haftmann
parents: 32345
diff changeset
    38
    val raw_thms = Code.these_eqns thy c'
0ac26087464b moved all technical processing of code equations to code_thingol.ML
haftmann
parents: 32345
diff changeset
    39
      |> map_filter (fn (thm, linear) => if linear then SOME thm else NONE);
0ac26087464b moved all technical processing of code equations to code_thingol.ML
haftmann
parents: 32345
diff changeset
    40
  in if null raw_thms then NONE else
0ac26087464b moved all technical processing of code equations to code_thingol.ML
haftmann
parents: 32345
diff changeset
    41
    raw_thms
0ac26087464b moved all technical processing of code equations to code_thingol.ML
haftmann
parents: 32345
diff changeset
    42
    |> Code_Thingol.clean_thms thy (snd (Code.const_typ_eqn thy (hd raw_thms)))
0ac26087464b moved all technical processing of code equations to code_thingol.ML
haftmann
parents: 32345
diff changeset
    43
    |> map (rpair opt_name)
0ac26087464b moved all technical processing of code equations to code_thingol.ML
haftmann
parents: 32345
diff changeset
    44
    |> SOME
0ac26087464b moved all technical processing of code equations to code_thingol.ML
haftmann
parents: 32345
diff changeset
    45
  end;
28522
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 28370
diff changeset
    46
31954
8db19c99b00a Stefan Berghofer's code generator uses Pure equality instead of HOL equality now
haftmann
parents: 31156
diff changeset
    47
val dest_eqn = Logic.dest_equals;
28522
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 28370
diff changeset
    48
val const_of = dest_Const o head_of o fst o dest_eqn;
15321
694f9d3ce90d Reimplemented some operations on "code lemma" table to avoid that code
berghofe
parents: 15257
diff changeset
    49
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    50
fun get_equations thy defs (s, T) =
28522
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 28370
diff changeset
    51
  (case retrieve_equations thy (s, T) of
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    52
     NONE => ([], "")
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    53
   | SOME thms => 
28522
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 28370
diff changeset
    54
       let val thms' = filter (fn (thm, _) => is_instance T
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 28370
diff changeset
    55
           (snd (const_of (prop_of thm)))) thms
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    56
       in if null thms' then ([], "")
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    57
         else (preprocess thy (map fst thms'),
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    58
           case snd (snd (split_last thms')) of
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    59
               NONE => (case get_defn thy defs s T of
27398
768da1da59d6 simplified retrieval of theory names of consts and types
haftmann
parents: 26975
diff changeset
    60
                   NONE => Codegen.thyname_of_const thy s
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    61
                 | SOME ((_, (thyname, _)), _) => thyname)
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    62
             | SOME thyname => thyname)
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    63
       end);
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    64
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    65
fun mk_suffix thy defs (s, T) = (case get_defn thy defs s T of
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
    66
  SOME (_, SOME i) => " def" ^ string_of_int i | _ => "");
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    67
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    68
exception EQN of string * typ * string;
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    69
22887
haftmann
parents: 22846
diff changeset
    70
fun cycle g (xs, x : string) =
haftmann
parents: 22846
diff changeset
    71
  if member (op =) xs x then xs
haftmann
parents: 22846
diff changeset
    72
  else Library.foldl (cycle g) (x :: xs, flat (Graph.all_paths (fst g) (x, x)));
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    73
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    74
fun add_rec_funs thy defs dep module eqs gr =
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    75
  let
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    76
    fun dest_eq t = (fst (const_of t) ^ mk_suffix thy defs (const_of t),
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
    77
      dest_eqn (rename_term t));
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    78
    val eqs' = map dest_eq eqs;
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    79
    val (dname, _) :: _ = eqs';
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    80
    val (s, T) = const_of (hd eqs);
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    81
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    82
    fun mk_fundef module fname first [] gr = ([], gr)
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    83
      | mk_fundef module fname first ((fname' : string, (lhs, rhs)) :: xs) gr =
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    84
      let
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    85
        val (pl, gr1) = invoke_codegen thy defs dname module false lhs gr;
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    86
        val (pr, gr2) = invoke_codegen thy defs dname module false rhs gr1;
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    87
        val (rest, gr3) = mk_fundef module fname' false xs gr2 ;
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    88
        val (ty, gr4) = invoke_tycodegen thy defs dname module false T gr3;
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    89
        val num_args = (length o snd o strip_comb) lhs;
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    90
        val prfx = if fname = fname' then "  |"
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    91
          else if not first then "and"
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    92
          else if num_args = 0 then "val"
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    93
          else "fun";
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    94
        val pl' = Pretty.breaks (str prfx
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    95
          :: (if num_args = 0 then [pl, str ":", ty] else [pl]));
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    96
      in
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    97
        (Pretty.blk (4, pl'
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
    98
           @ [str " =", Pretty.brk 1, pr]) :: rest, gr4)
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
    99
      end;
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   100
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   101
    fun put_code module fundef = map_node dname
26975
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26928
diff changeset
   102
      (K (SOME (EQN ("", dummyT, dname)), module, string_of (Pretty.blk (0,
103dca19ef2e Replaced Pretty.str and Pretty.string_of by specific functions (from Codegen) that
berghofe
parents: 26928
diff changeset
   103
      separate Pretty.fbrk fundef @ [str ";"])) ^ "\n\n"));
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   104
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   105
  in
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   106
    (case try (get_node gr) dname of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15321
diff changeset
   107
       NONE =>
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   108
         let
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   109
           val gr1 = add_edge (dname, dep)
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   110
             (new_node (dname, (SOME (EQN (s, T, "")), module, "")) gr);
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   111
           val (fundef, gr2) = mk_fundef module "" true eqs' gr1 ;
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   112
           val xs = cycle gr2 ([], dname);
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   113
           val cs = map (fn x => case get_node gr2 x of
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
   114
               (SOME (EQN (s, T, _)), _, _) => (s, T)
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   115
             | _ => error ("RecfunCodegen: illegal cyclic dependencies:\n" ^
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   116
                implode (separate ", " xs))) xs
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   117
         in (case xs of
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   118
             [_] => (module, put_code module fundef gr2)
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   119
           | _ =>
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   120
             if not (dep mem xs) then
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   121
               let
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
   122
                 val thmss as (_, thyname) :: _ = map (get_equations thy defs) cs;
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   123
                 val module' = if_library thyname module;
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
   124
                 val eqs'' = map (dest_eq o prop_of) (List.concat (map fst thmss));
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   125
                 val (fundef', gr3) = mk_fundef module' "" true eqs''
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   126
                   (add_edge (dname, dep)
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 29272
diff changeset
   127
                     (List.foldr (uncurry new_node) (del_nodes xs gr2)
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   128
                       (map (fn k =>
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   129
                         (k, (SOME (EQN ("", dummyT, dname)), module', ""))) xs)))
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   130
               in (module', put_code module' fundef' gr3) end
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   131
             else (module, gr2))
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   132
         end
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   133
     | SOME (SOME (EQN (_, _, s)), module', _) =>
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   134
         (module', if s = "" then
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   135
            if dname = dep then gr else add_edge (dname, dep) gr
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   136
          else if s = dep then gr else add_edge (s, dep) gr))
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   137
  end;
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   138
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   139
fun recfun_codegen thy defs dep module brack t gr = (case strip_comb t of
22921
475ff421a6a3 consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents: 22887
diff changeset
   140
    (Const (p as (s, T)), ts) => (case (get_equations thy defs p, get_assoc_code thy (s, T)) of
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
   141
       (([], _), _) => NONE
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15321
diff changeset
   142
     | (_, SOME _) => NONE
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   143
     | ((eqns, thyname), NONE) =>
16645
a152d6b21c31 Adapted to modular code generation.
berghofe
parents: 16424
diff changeset
   144
        let
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   145
          val module' = if_library thyname module;
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   146
          val (ps, gr') = fold_map
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   147
            (invoke_codegen thy defs dep module true) ts gr;
17144
6642e0f96f44 Implemented incremental code generation.
berghofe
parents: 16645
diff changeset
   148
          val suffix = mk_suffix thy defs p;
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   149
          val (module'', gr'') =
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   150
            add_rec_funs thy defs dep module' (map prop_of eqns) gr';
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   151
          val (fname, gr''') = mk_const_id module'' (s ^ suffix) gr''
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   152
        in
28535
38fb0780b58b made SMLNJ happy
haftmann
parents: 28522
diff changeset
   153
          SOME (mk_app brack (str (mk_qual_id module fname)) ps, gr''')
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   154
        end)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15321
diff changeset
   155
  | _ => NONE);
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   156
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31962
diff changeset
   157
val setup = 
24219
e558fe311376 new structure for code generator modules
haftmann
parents: 22921
diff changeset
   158
  add_codegen "recfun" recfun_codegen
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31962
diff changeset
   159
  #> Code.set_code_target_attr add_thm_target;
12447
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   160
e752c9aecdec Code generator for recursive functions.
berghofe
parents:
diff changeset
   161
end;