| 
6429
 | 
     1  | 
(*  Title:      HOL/Tools/recdef_package.ML
  | 
| 
 | 
     2  | 
    ID:         $Id$
  | 
| 
 | 
     3  | 
    Author:     Markus Wenzel, TU Muenchen
  | 
| 
 | 
     4  | 
  | 
| 
 | 
     5  | 
Wrapper module for Konrad Slind's TFL package.
  | 
| 
 | 
     6  | 
*)
  | 
| 
 | 
     7  | 
  | 
| 
 | 
     8  | 
signature RECDEF_PACKAGE =
  | 
| 
 | 
     9  | 
sig
  | 
| 
 | 
    10  | 
  val quiet_mode: bool ref
  | 
| 
6439
 | 
    11  | 
  val print_recdefs: theory -> unit
  | 
| 
 | 
    12  | 
  val get_recdef: theory -> string -> {rules: thm list, induct: thm, tcs: term list}
 | 
| 
6429
 | 
    13  | 
  val add_recdef: xstring -> string -> ((bstring * string) * Args.src list) list
  | 
| 
 | 
    14  | 
    -> string option -> (xstring * Args.src list) list
  | 
| 
 | 
    15  | 
    -> theory -> theory * {rules: thm list, induct: thm, tcs: term list}
 | 
| 
 | 
    16  | 
  val add_recdef_i: xstring -> term -> ((bstring * term) * theory attribute list) list
  | 
| 
 | 
    17  | 
    -> simpset option -> (thm * theory attribute list) list
  | 
| 
 | 
    18  | 
    -> theory -> theory * {rules: thm list, induct: thm, tcs: term list}
 | 
| 
6439
 | 
    19  | 
  val setup: (theory -> theory) list
  | 
| 
6429
 | 
    20  | 
end;
  | 
| 
 | 
    21  | 
  | 
| 
 | 
    22  | 
structure RecdefPackage: RECDEF_PACKAGE =
  | 
| 
 | 
    23  | 
struct
  | 
| 
 | 
    24  | 
  | 
| 
6458
 | 
    25  | 
(* FIXME tmp
  | 
| 
6429
 | 
    26  | 
val quiet_mode = Tfl.quiet_mode;
  | 
| 
 | 
    27  | 
val message = Tfl.message;
  | 
| 
6458
 | 
    28  | 
*)
  | 
| 
 | 
    29  | 
val quiet_mode = ref false;
  | 
| 
 | 
    30  | 
val message = writeln;
  | 
| 
6429
 | 
    31  | 
  | 
| 
 | 
    32  | 
  | 
| 
6439
 | 
    33  | 
  | 
| 
 | 
    34  | 
(** theory data **)
  | 
| 
 | 
    35  | 
  | 
| 
 | 
    36  | 
(* data kind 'HOL/recdef' *)
  | 
| 
 | 
    37  | 
  | 
| 
 | 
    38  | 
type recdef_info = {rules: thm list, induct: thm, tcs: term list};
 | 
| 
 | 
    39  | 
  | 
| 
 | 
    40  | 
structure RecdefArgs =
  | 
| 
 | 
    41  | 
struct
  | 
| 
 | 
    42  | 
  val name = "HOL/recdef";
  | 
| 
 | 
    43  | 
  type T = recdef_info Symtab.table;
  | 
| 
 | 
    44  | 
  | 
| 
 | 
    45  | 
  val empty = Symtab.empty;
  | 
| 
 | 
    46  | 
  val prep_ext = I;
  | 
| 
 | 
    47  | 
  val merge: T * T -> T = Symtab.merge (K true);
  | 
| 
 | 
    48  | 
  | 
| 
 | 
    49  | 
  fun print sg tab =
  | 
| 
 | 
    50  | 
    Pretty.writeln (Pretty.strs ("recdefs:" ::
 | 
| 
 | 
    51  | 
      map (Sign.cond_extern sg Sign.constK o fst) (Symtab.dest tab)));
  | 
| 
 | 
    52  | 
end;
  | 
| 
 | 
    53  | 
  | 
| 
 | 
    54  | 
structure RecdefData = TheoryDataFun(RecdefArgs);
  | 
| 
 | 
    55  | 
val print_recdefs = RecdefData.print;
  | 
| 
6429
 | 
    56  | 
  | 
| 
6439
 | 
    57  | 
  | 
| 
 | 
    58  | 
(* get and put data *)
  | 
| 
 | 
    59  | 
  | 
| 
 | 
    60  | 
fun get_recdef thy name =
  | 
| 
 | 
    61  | 
  (case Symtab.lookup (RecdefData.get thy, name) of
  | 
| 
 | 
    62  | 
    Some info => info
  | 
| 
 | 
    63  | 
  | None => error ("Unknown recursive function " ^ quote name));
 | 
| 
 | 
    64  | 
  | 
| 
 | 
    65  | 
fun put_recdef name info thy =
  | 
| 
6429
 | 
    66  | 
  let
  | 
| 
6439
 | 
    67  | 
    val tab = Symtab.update_new ((name, info), RecdefData.get thy)
  | 
| 
 | 
    68  | 
      handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name);
 | 
| 
 | 
    69  | 
  in RecdefData.put tab thy end;
  | 
| 
 | 
    70  | 
  | 
| 
 | 
    71  | 
  | 
| 
 | 
    72  | 
  | 
| 
 | 
    73  | 
(** add_recdef(_i) **)
  | 
| 
 | 
    74  | 
  | 
| 
 | 
    75  | 
fun gen_add_recdef tfl_def prep_att prep_ss app_thms raw_name R eq_srcs raw_ss raw_congs thy =
  | 
| 
 | 
    76  | 
  let
  | 
| 
 | 
    77  | 
    val name = Sign.intern_const (Theory.sign_of thy) raw_name;
  | 
| 
 | 
    78  | 
    val bname = Sign.base_name name;
  | 
| 
 | 
    79  | 
  | 
| 
6429
 | 
    80  | 
    val _ = message ("Defining recursive function " ^ quote name ^ " ...");
 | 
| 
 | 
    81  | 
  | 
| 
 | 
    82  | 
    val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs);
  | 
| 
 | 
    83  | 
    val eq_atts = map (map (prep_att thy)) raw_eq_atts;
  | 
| 
 | 
    84  | 
    val ss = (case raw_ss of None => Simplifier.simpset_of thy | Some x => prep_ss x);
  | 
| 
 | 
    85  | 
    val (thy1, congs) = thy |> app_thms raw_congs;
  | 
| 
 | 
    86  | 
    val (thy2, pats) = tfl_def thy1 name R eqs;
  | 
| 
 | 
    87  | 
    val (result as {rules, induct, tcs}) = Tfl.simplify_defn (ss, congs) (thy2, (name, pats));
 | 
| 
 | 
    88  | 
    val thy3 =
  | 
| 
 | 
    89  | 
      thy2
  | 
| 
6439
 | 
    90  | 
      |> Theory.add_path bname
  | 
| 
6429
 | 
    91  | 
      |> PureThy.add_thmss [(("rules", rules), [])]
 | 
| 
 | 
    92  | 
      |> PureThy.add_thms ((("induct", induct), []) :: ((eq_names ~~ rules) ~~ eq_atts))
 | 
| 
6439
 | 
    93  | 
      |> put_recdef name result
  | 
| 
6429
 | 
    94  | 
      |> Theory.parent_path;
  | 
| 
 | 
    95  | 
  in (thy3, result) end;
  | 
| 
 | 
    96  | 
  | 
| 
 | 
    97  | 
val add_recdef = gen_add_recdef Tfl.define Attrib.global_attribute
  | 
| 
 | 
    98  | 
  (Simplifier.simpset_of o ThyInfo.get_theory) IsarThy.apply_theorems;
  | 
| 
 | 
    99  | 
  | 
| 
 | 
   100  | 
val add_recdef_i = gen_add_recdef Tfl.define_i (K I) I IsarThy.apply_theorems_i;
  | 
| 
 | 
   101  | 
  | 
| 
 | 
   102  | 
  | 
| 
6439
 | 
   103  | 
(** package setup **)
  | 
| 
 | 
   104  | 
  | 
| 
 | 
   105  | 
(* setup theory *)
  | 
| 
 | 
   106  | 
  | 
| 
 | 
   107  | 
val setup = [RecdefData.init];
  | 
| 
 | 
   108  | 
  | 
| 
 | 
   109  | 
  | 
| 
6429
 | 
   110  | 
(* outer syntax *)
  | 
| 
 | 
   111  | 
  | 
| 
 | 
   112  | 
local open OuterParse in
  | 
| 
 | 
   113  | 
  | 
| 
 | 
   114  | 
val recdef_decl =
  | 
| 
 | 
   115  | 
  name -- term -- Scan.repeat1 (opt_thm_name ":" -- term) --
  | 
| 
 | 
   116  | 
  Scan.optional ($$$ "congs" |-- !!! xthms1) [] --
  | 
| 
 | 
   117  | 
  Scan.option ($$$ "simpset" |-- !!! name)
  | 
| 
 | 
   118  | 
  >> (fn ((((f, R), eqs), congs), ss) => #1 o add_recdef f R (map triple_swap eqs) ss congs);
  | 
| 
 | 
   119  | 
  | 
| 
 | 
   120  | 
val recdefP =
  | 
| 
6439
 | 
   121  | 
  OuterSyntax.command "recdef" "define general recursive functions (TFL)"
  | 
| 
6429
 | 
   122  | 
    (recdef_decl >> Toplevel.theory);
  | 
| 
 | 
   123  | 
  | 
| 
 | 
   124  | 
val _ = OuterSyntax.add_keywords ["congs", "simpset"];
  | 
| 
 | 
   125  | 
val _ = OuterSyntax.add_parsers [recdefP];
  | 
| 
 | 
   126  | 
  | 
| 
 | 
   127  | 
end;
  | 
| 
 | 
   128  | 
  | 
| 
 | 
   129  | 
  | 
| 
 | 
   130  | 
end;
  |