src/HOL/Tools/recfun_codegen.ML
author haftmann
Sat, 15 Sep 2007 19:27:35 +0200
changeset 24584 01e83ffa6c54
parent 24219 e558fe311376
child 24624 b8383b1bbae3
permissions -rw-r--r--
fixed title
     1 (*  Title:      HOL/Tools/recfun_codegen.ML
     2     ID:         $Id$
     3     Author:     Stefan Berghofer, TU Muenchen
     4 
     5 Code generator for recursive functions.
     6 *)
     7 
     8 signature RECFUN_CODEGEN =
     9 sig
    10   val add: string option -> attribute
    11   val del: attribute
    12   val setup: theory -> theory
    13 end;
    14 
    15 structure RecfunCodegen : RECFUN_CODEGEN =
    16 struct
    17 
    18 open Codegen;
    19 
    20 structure RecCodegenData = TheoryDataFun
    21 (
    22   type T = (thm * string option) list Symtab.table;
    23   val empty = Symtab.empty;
    24   val copy = I;
    25   val extend = I;
    26   fun merge _ = Symtab.merge_list (Thm.eq_thm_prop o pairself fst);
    27 );
    28 
    29 
    30 val dest_eqn = HOLogic.dest_eq o HOLogic.dest_Trueprop;
    31 val lhs_of = fst o dest_eqn o prop_of;
    32 val const_of = dest_Const o head_of o fst o dest_eqn;
    33 
    34 fun warn thm = warning ("RecfunCodegen: Not a proper equation:\n" ^
    35   string_of_thm thm);
    36 
    37 fun add optmod =
    38   let
    39     fun add thm =
    40       if Pattern.pattern (lhs_of thm) then
    41         RecCodegenData.map
    42           (Symtab.update_list ((fst o const_of o prop_of) thm, (thm, optmod)))
    43       else tap (fn _ => warn thm)
    44       handle TERM _ => tap (fn _ => warn thm);
    45   in
    46     Thm.declaration_attribute (fn thm => Context.mapping
    47       (add thm #> Code.add_func false thm) I)
    48   end;
    49 
    50 fun del_thm thm thy =
    51   let
    52     val tab = RecCodegenData.get thy;
    53     val (s, _) = const_of (prop_of thm);
    54   in case Symtab.lookup tab s of
    55       NONE => thy
    56     | SOME thms =>
    57         RecCodegenData.put (Symtab.update (s, remove (Thm.eq_thm o apsnd fst) thm thms) tab) thy
    58   end handle TERM _ => (warn thm; thy);
    59 
    60 val del = Thm.declaration_attribute
    61   (fn thm => Context.mapping (del_thm thm #> Code.del_func thm) I)
    62 
    63 fun del_redundant thy eqs [] = eqs
    64   | del_redundant thy eqs (eq :: eqs') =
    65     let
    66       val matches = curry
    67         (Pattern.matches thy o pairself (lhs_of o fst))
    68     in del_redundant thy (eq :: eqs) (filter_out (matches eq) eqs') end;
    69 
    70 fun get_equations thy defs (s, T) =
    71   (case Symtab.lookup (RecCodegenData.get thy) s of
    72      NONE => ([], "")
    73    | SOME thms => 
    74        let val thms' = del_redundant thy []
    75          (filter (fn (thm, _) => is_instance thy T
    76            (snd (const_of (prop_of thm)))) thms)
    77        in if null thms' then ([], "")
    78          else (preprocess thy (map fst thms'),
    79            case snd (snd (split_last thms')) of
    80                NONE => (case get_defn thy defs s T of
    81                    NONE => thyname_of_const s thy
    82                  | SOME ((_, (thyname, _)), _) => thyname)
    83              | SOME thyname => thyname)
    84        end);
    85 
    86 fun mk_suffix thy defs (s, T) = (case get_defn thy defs s T of
    87   SOME (_, SOME i) => " def" ^ string_of_int i | _ => "");
    88 
    89 exception EQN of string * typ * string;
    90 
    91 fun cycle g (xs, x : string) =
    92   if member (op =) xs x then xs
    93   else Library.foldl (cycle g) (x :: xs, flat (Graph.all_paths (fst g) (x, x)));
    94 
    95 fun add_rec_funs thy defs gr dep eqs module =
    96   let
    97     fun dest_eq t = (fst (const_of t) ^ mk_suffix thy defs (const_of t),
    98       dest_eqn (rename_term t));
    99     val eqs' = map dest_eq eqs;
   100     val (dname, _) :: _ = eqs';
   101     val (s, T) = const_of (hd eqs);
   102 
   103     fun mk_fundef module fname prfx gr [] = (gr, [])
   104       | mk_fundef module fname prfx gr ((fname' : string, (lhs, rhs)) :: xs) =
   105       let
   106         val (gr1, pl) = invoke_codegen thy defs dname module false (gr, lhs);
   107         val (gr2, pr) = invoke_codegen thy defs dname module false (gr1, rhs);
   108         val (gr3, rest) = mk_fundef module fname' "and " gr2 xs
   109       in
   110         (gr3, Pretty.blk (4, [Pretty.str (if fname = fname' then "  | " else prfx),
   111            pl, Pretty.str " =", Pretty.brk 1, pr]) :: rest)
   112       end;
   113 
   114     fun put_code module fundef = map_node dname
   115       (K (SOME (EQN ("", dummyT, dname)), module, Pretty.string_of (Pretty.blk (0,
   116       separate Pretty.fbrk fundef @ [Pretty.str ";"])) ^ "\n\n"));
   117 
   118   in
   119     (case try (get_node gr) dname of
   120        NONE =>
   121          let
   122            val gr1 = add_edge (dname, dep)
   123              (new_node (dname, (SOME (EQN (s, T, "")), module, "")) gr);
   124            val (gr2, fundef) = mk_fundef module "" "fun " gr1 eqs';
   125            val xs = cycle gr2 ([], dname);
   126            val cs = map (fn x => case get_node gr2 x of
   127                (SOME (EQN (s, T, _)), _, _) => (s, T)
   128              | _ => error ("RecfunCodegen: illegal cyclic dependencies:\n" ^
   129                 implode (separate ", " xs))) xs
   130          in (case xs of
   131              [_] => (put_code module fundef gr2, module)
   132            | _ =>
   133              if not (dep mem xs) then
   134                let
   135                  val thmss as (_, thyname) :: _ = map (get_equations thy defs) cs;
   136                  val module' = if_library thyname module;
   137                  val eqs'' = map (dest_eq o prop_of) (List.concat (map fst thmss));
   138                  val (gr3, fundef') = mk_fundef module' "" "fun "
   139                    (add_edge (dname, dep)
   140                      (foldr (uncurry new_node) (del_nodes xs gr2)
   141                        (map (fn k =>
   142                          (k, (SOME (EQN ("", dummyT, dname)), module', ""))) xs))) eqs''
   143                in (put_code module' fundef' gr3, module') end
   144              else (gr2, module))
   145          end
   146      | SOME (SOME (EQN (_, _, s)), module', _) =>
   147          (if s = "" then
   148             if dname = dep then gr else add_edge (dname, dep) gr
   149           else if s = dep then gr else add_edge (s, dep) gr,
   150           module'))
   151   end;
   152 
   153 fun recfun_codegen thy defs gr dep module brack t = (case strip_comb t of
   154     (Const (p as (s, T)), ts) => (case (get_equations thy defs p, get_assoc_code thy (s, T)) of
   155        (([], _), _) => NONE
   156      | (_, SOME _) => NONE
   157      | ((eqns, thyname), NONE) =>
   158         let
   159           val module' = if_library thyname module;
   160           val (gr', ps) = foldl_map
   161             (invoke_codegen thy defs dep module true) (gr, ts);
   162           val suffix = mk_suffix thy defs p;
   163           val (gr'', module'') =
   164             add_rec_funs thy defs gr' dep (map prop_of eqns) module';
   165           val (gr''', fname) = mk_const_id module'' (s ^ suffix) gr''
   166         in
   167           SOME (gr''', mk_app brack (Pretty.str (mk_qual_id module fname)) ps)
   168         end)
   169   | _ => NONE);
   170 
   171 
   172 val setup =
   173   add_codegen "recfun" recfun_codegen
   174   #> Code.add_attribute ("", Args.del |-- Scan.succeed del
   175      || Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) >> add);
   176 
   177 end;