src/Pure/Tools/named_thms.ML
author berghofe
Tue, 27 Nov 2007 15:44:49 +0100
changeset 25471 ca009b7ce693
parent 24867 e5b55d7be9bb
child 26363 9d95309f8069
permissions -rw-r--r--
Optimized beta_norm: only tries to normalize term when it contains abstractions.

(*  Title:      Pure/Tools/named_thms.ML
    ID:         $Id$
    Author:     Makarius

Named collections of theorems in canonical order.
*)

signature NAMED_THMS =
sig
  val get: Proof.context -> thm list
  val pretty: Proof.context -> Pretty.T
  val add: attribute
  val del: attribute
  val setup: theory -> theory
end;

functor NamedThmsFun(val name: string val description: string): NAMED_THMS =
struct

structure Data = GenericDataFun
(
  type T = thm list;
  val empty = [];
  val extend = I;
  fun merge _ = Thm.merge_thms;
);

val get = Data.get o Context.Proof;

fun pretty ctxt =
  Pretty.big_list (description ^ ":") (map (ProofContext.pretty_thm ctxt) (get ctxt));

val add = Thm.declaration_attribute (Data.map o Thm.add_thm);
val del = Thm.declaration_attribute (Data.map o Thm.del_thm);

val setup =
  Attrib.add_attributes [(name, Attrib.add_del_args add del, "declaration of " ^ description)];

val _ =
  OuterSyntax.improper_command ("print_" ^ name ^ "_rules") ("print " ^ description)
    OuterKeyword.diag
    (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o
      Toplevel.keep (Pretty.writeln o pretty o Toplevel.context_of)));

end;