src/HOL/Tools/recfun_codegen.ML
author haftmann
Tue, 17 Jan 2006 16:36:57 +0100
changeset 18702 7dc7dcd63224
parent 18220 43cf5767f992
child 18708 4b3dadb4fe33
permissions -rw-r--r--
substantial improvements in code generator
berghofe@12447
     1
(*  Title:      HOL/recfun_codegen.ML
berghofe@12447
     2
    ID:         $Id$
berghofe@12447
     3
    Author:     Stefan Berghofer, TU Muenchen
berghofe@12447
     4
berghofe@12447
     5
Code generator for recursive functions.
berghofe@12447
     6
*)
berghofe@12447
     7
berghofe@12447
     8
signature RECFUN_CODEGEN =
berghofe@12447
     9
sig
berghofe@16645
    10
  val add: string option -> theory attribute
berghofe@16645
    11
  val del: theory attribute
haftmann@18220
    12
  val get_rec_equations: theory -> string * typ -> (term list * term) list * typ
haftmann@18702
    13
  val get_thm_equations: theory -> string * typ -> (thm list * typ) option
berghofe@12447
    14
  val setup: (theory -> theory) list
berghofe@12447
    15
end;
berghofe@12447
    16
berghofe@12447
    17
structure RecfunCodegen : RECFUN_CODEGEN =
berghofe@12447
    18
struct
berghofe@12447
    19
berghofe@12447
    20
open Codegen;
berghofe@12447
    21
wenzelm@16424
    22
structure CodegenData = TheoryDataFun
wenzelm@16424
    23
(struct
berghofe@12447
    24
  val name = "HOL/recfun_codegen";
berghofe@16645
    25
  type T = (thm * string option) list Symtab.table;
berghofe@12447
    26
  val empty = Symtab.empty;
berghofe@12447
    27
  val copy = I;
wenzelm@16424
    28
  val extend = I;
berghofe@16645
    29
  fun merge _ = Symtab.merge_multi' (Drule.eq_thm_prop o pairself fst);
berghofe@12447
    30
  fun print _ _ = ();
wenzelm@16424
    31
end);
berghofe@12447
    32
berghofe@12447
    33
berghofe@12447
    34
val dest_eqn = HOLogic.dest_eq o HOLogic.dest_Trueprop;
berghofe@15321
    35
val lhs_of = fst o dest_eqn o prop_of;
berghofe@12447
    36
val const_of = dest_Const o head_of o fst o dest_eqn;
berghofe@12447
    37
berghofe@12447
    38
fun warn thm = warning ("RecfunCodegen: Not a proper equation:\n" ^
berghofe@12447
    39
  string_of_thm thm);
berghofe@12447
    40
berghofe@16645
    41
fun add optmod (p as (thy, thm)) =
berghofe@12447
    42
  let
berghofe@12556
    43
    val tab = CodegenData.get thy;
berghofe@12447
    44
    val (s, _) = const_of (prop_of thm);
berghofe@15321
    45
  in
berghofe@15321
    46
    if Pattern.pattern (lhs_of thm) then
wenzelm@17412
    47
      (CodegenData.put (Symtab.update_multi (s, (thm, optmod)) tab) thy, thm)
berghofe@15321
    48
    else (warn thm; p)
berghofe@15321
    49
  end handle TERM _ => (warn thm; p);
berghofe@12447
    50
berghofe@14196
    51
fun del (p as (thy, thm)) =
berghofe@14196
    52
  let
berghofe@14196
    53
    val tab = CodegenData.get thy;
berghofe@14196
    54
    val (s, _) = const_of (prop_of thm);
wenzelm@17412
    55
  in case Symtab.lookup tab s of
skalberg@15531
    56
      NONE => p
wenzelm@17412
    57
    | SOME thms => (CodegenData.put (Symtab.update (s,
wenzelm@17261
    58
        gen_rem (eq_thm o apfst fst) (thms, thm)) tab) thy, thm)
berghofe@14196
    59
  end handle TERM _ => (warn thm; p);
berghofe@14196
    60
berghofe@15321
    61
fun del_redundant thy eqs [] = eqs
berghofe@15321
    62
  | del_redundant thy eqs (eq :: eqs') =
berghofe@15321
    63
    let
berghofe@15321
    64
      val matches = curry
wenzelm@17203
    65
        (Pattern.matches thy o pairself (lhs_of o fst))
berghofe@15321
    66
    in del_redundant thy (eq :: eqs) (filter_out (matches eq) eqs') end;
berghofe@15321
    67
berghofe@16645
    68
fun get_equations thy defs (s, T) =
wenzelm@17412
    69
  (case Symtab.lookup (CodegenData.get thy) s of
berghofe@16645
    70
     NONE => ([], "")
berghofe@16645
    71
   | SOME thms => 
berghofe@16645
    72
       let val thms' = del_redundant thy []
berghofe@16645
    73
         (List.filter (fn (thm, _) => is_instance thy T
berghofe@16645
    74
           (snd (const_of (prop_of thm)))) thms)
berghofe@16645
    75
       in if null thms' then ([], "")
berghofe@16645
    76
         else (preprocess thy (map fst thms'),
berghofe@16645
    77
           case snd (snd (split_last thms')) of
berghofe@16645
    78
               NONE => (case get_defn thy defs s T of
berghofe@16645
    79
                   NONE => thyname_of_const s thy
berghofe@16645
    80
                 | SOME ((_, (thyname, _)), _) => thyname)
berghofe@16645
    81
             | SOME thyname => thyname)
berghofe@16645
    82
       end);
berghofe@12447
    83
haftmann@18220
    84
fun get_rec_equations thy (s, T) =
haftmann@18220
    85
  Symtab.lookup (CodegenData.get thy) s
haftmann@18220
    86
  |> Option.map (fn thms => 
haftmann@18220
    87
       List.filter (fn (thm, _) => is_instance thy T ((snd o const_of o prop_of) thm)) thms
haftmann@18220
    88
       |> del_redundant thy [])
haftmann@18220
    89
  |> Option.mapPartial (fn thms => if null thms then NONE else SOME thms)
haftmann@18220
    90
  |> Option.map (fn thms =>
haftmann@18220
    91
       (preprocess thy (map fst thms),
haftmann@18220
    92
          (snd o const_of o prop_of o fst o hd) thms))
haftmann@18220
    93
  |> the_default ([], dummyT)
haftmann@18220
    94
  |> apfst (map prop_of)
haftmann@18220
    95
  |> apfst (map (Codegen.rename_term #> HOLogic.dest_Trueprop #> HOLogic.dest_eq #> apfst (strip_comb #> snd)))
haftmann@18220
    96
haftmann@18702
    97
fun get_thm_equations thy (c, ty) =
haftmann@18702
    98
  Symtab.lookup (CodegenData.get thy) c
haftmann@18702
    99
  |> Option.map (fn thms => 
haftmann@18702
   100
       List.filter (fn (thm, _) => is_instance thy ty ((snd o const_of o prop_of) thm)) thms
haftmann@18702
   101
       |> del_redundant thy [])
haftmann@18702
   102
  |> Option.mapPartial (fn thms => if null thms then NONE else SOME thms)
haftmann@18702
   103
  |> Option.map (fn thms =>
haftmann@18702
   104
       (preprocess thy (map fst thms),
haftmann@18702
   105
          (snd o const_of o prop_of o fst o hd) thms))
haftmann@18702
   106
  |> (Option.map o apfst o map) (fn thm => thm RS HOL.eq_reflection);
haftmann@18702
   107
berghofe@16645
   108
fun mk_suffix thy defs (s, T) = (case get_defn thy defs s T of
berghofe@17144
   109
  SOME (_, SOME i) => " def" ^ string_of_int i | _ => "");
berghofe@12447
   110
berghofe@12447
   111
exception EQN of string * typ * string;
berghofe@12447
   112
berghofe@12447
   113
fun cycle g (xs, x) =
berghofe@12447
   114
  if x mem xs then xs
berghofe@17144
   115
  else Library.foldl (cycle g) (x :: xs, List.concat (Graph.find_paths (fst g) (x, x)));
berghofe@12447
   116
berghofe@17144
   117
fun add_rec_funs thy defs gr dep eqs module =
berghofe@12447
   118
  let
berghofe@16645
   119
    fun dest_eq t = (fst (const_of t) ^ mk_suffix thy defs (const_of t),
berghofe@16645
   120
      dest_eqn (rename_term t));
berghofe@12447
   121
    val eqs' = map dest_eq eqs;
berghofe@12447
   122
    val (dname, _) :: _ = eqs';
berghofe@12447
   123
    val (s, T) = const_of (hd eqs);
berghofe@12447
   124
berghofe@17144
   125
    fun mk_fundef module fname prfx gr [] = (gr, [])
berghofe@17144
   126
      | mk_fundef module fname prfx gr ((fname', (lhs, rhs)) :: xs) =
berghofe@12447
   127
      let
berghofe@17144
   128
        val (gr1, pl) = invoke_codegen thy defs dname module false (gr, lhs);
berghofe@17144
   129
        val (gr2, pr) = invoke_codegen thy defs dname module false (gr1, rhs);
berghofe@17144
   130
        val (gr3, rest) = mk_fundef module fname' "and " gr2 xs
berghofe@12447
   131
      in
berghofe@12447
   132
        (gr3, Pretty.blk (4, [Pretty.str (if fname = fname' then "  | " else prfx),
berghofe@12447
   133
           pl, Pretty.str " =", Pretty.brk 1, pr]) :: rest)
berghofe@12447
   134
      end;
berghofe@12447
   135
berghofe@17144
   136
    fun put_code module fundef = map_node dname
berghofe@17144
   137
      (K (SOME (EQN ("", dummyT, dname)), module, Pretty.string_of (Pretty.blk (0,
berghofe@12447
   138
      separate Pretty.fbrk fundef @ [Pretty.str ";"])) ^ "\n\n"));
berghofe@12447
   139
berghofe@12447
   140
  in
berghofe@17144
   141
    (case try (get_node gr) dname of
skalberg@15531
   142
       NONE =>
berghofe@12447
   143
         let
berghofe@17144
   144
           val gr1 = add_edge (dname, dep)
berghofe@17144
   145
             (new_node (dname, (SOME (EQN (s, T, "")), module, "")) gr);
berghofe@17144
   146
           val (gr2, fundef) = mk_fundef module "" "fun " gr1 eqs';
berghofe@12447
   147
           val xs = cycle gr2 ([], dname);
berghofe@17144
   148
           val cs = map (fn x => case get_node gr2 x of
berghofe@16645
   149
               (SOME (EQN (s, T, _)), _, _) => (s, T)
berghofe@12447
   150
             | _ => error ("RecfunCodegen: illegal cyclic dependencies:\n" ^
berghofe@12447
   151
                implode (separate ", " xs))) xs
berghofe@12447
   152
         in (case xs of
berghofe@17144
   153
             [_] => (put_code module fundef gr2, module)
berghofe@12447
   154
           | _ =>
berghofe@12447
   155
             if not (dep mem xs) then
berghofe@12447
   156
               let
berghofe@16645
   157
                 val thmss as (_, thyname) :: _ = map (get_equations thy defs) cs;
berghofe@17144
   158
                 val module' = if_library thyname module;
berghofe@16645
   159
                 val eqs'' = map (dest_eq o prop_of) (List.concat (map fst thmss));
berghofe@17144
   160
                 val (gr3, fundef') = mk_fundef module' "" "fun "
berghofe@17144
   161
                   (add_edge (dname, dep)
berghofe@17144
   162
                     (foldr (uncurry new_node) (del_nodes xs gr2)
skalberg@15574
   163
                       (map (fn k =>
berghofe@17144
   164
                         (k, (SOME (EQN ("", dummyT, dname)), module', ""))) xs))) eqs''
berghofe@17144
   165
               in (put_code module' fundef' gr3, module') end
berghofe@17144
   166
             else (gr2, module))
berghofe@12447
   167
         end
berghofe@17144
   168
     | SOME (SOME (EQN (_, _, s)), module', _) =>
berghofe@17144
   169
         (if s = "" then
berghofe@17144
   170
            if dname = dep then gr else add_edge (dname, dep) gr
berghofe@17144
   171
          else if s = dep then gr else add_edge (s, dep) gr,
berghofe@17144
   172
          module'))
berghofe@12447
   173
  end;
berghofe@12447
   174
berghofe@17144
   175
fun recfun_codegen thy defs gr dep module brack t = (case strip_comb t of
berghofe@16645
   176
    (Const (p as (s, T)), ts) => (case (get_equations thy defs p, get_assoc_code thy s T) of
berghofe@16645
   177
       (([], _), _) => NONE
skalberg@15531
   178
     | (_, SOME _) => NONE
berghofe@17144
   179
     | ((eqns, thyname), NONE) =>
berghofe@16645
   180
        let
berghofe@17144
   181
          val module' = if_library thyname module;
berghofe@16645
   182
          val (gr', ps) = foldl_map
berghofe@17144
   183
            (invoke_codegen thy defs dep module true) (gr, ts);
berghofe@17144
   184
          val suffix = mk_suffix thy defs p;
berghofe@17144
   185
          val (gr'', module'') =
berghofe@17144
   186
            add_rec_funs thy defs gr' dep (map prop_of eqns) module';
berghofe@17144
   187
          val (gr''', fname) = mk_const_id module'' (s ^ suffix) gr''
berghofe@12447
   188
        in
berghofe@17144
   189
          SOME (gr''', mk_app brack (Pretty.str (mk_qual_id module fname)) ps)
berghofe@12447
   190
        end)
skalberg@15531
   191
  | _ => NONE);
berghofe@12447
   192
berghofe@12447
   193
haftmann@18220
   194
val setup = [
haftmann@18220
   195
  CodegenData.init,
haftmann@18220
   196
  add_codegen "recfun" recfun_codegen,
haftmann@18220
   197
  add_attribute ""
haftmann@18220
   198
    (   Args.del |-- Scan.succeed del
haftmann@18220
   199
     || Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) >> add),
haftmann@18702
   200
  CodegenPackage.add_eqextr
haftmann@18702
   201
    ("rec", fn thy => fn _ => get_thm_equations thy)
haftmann@18220
   202
];
berghofe@12447
   203
berghofe@12447
   204
end;