src/Pure/Tools/named_thms.ML
author wenzelm
Sun, 15 Mar 2009 15:59:44 +0100
changeset 30528 7173bf123335
parent 29579 cb520b766e00
child 31901 e280491f36b8
permissions -rw-r--r--
simplified attribute setup;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24047
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Tools/named_thms.ML
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
     3
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
     4
Named collections of theorems in canonical order.
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
     5
*)
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
     6
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
     7
signature NAMED_THMS =
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
     8
sig
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
     9
  val get: Proof.context -> thm list
26363
9d95309f8069 export add/del_thm;
wenzelm
parents: 24867
diff changeset
    10
  val add_thm: thm -> Context.generic -> Context.generic
9d95309f8069 export add/del_thm;
wenzelm
parents: 24867
diff changeset
    11
  val del_thm: thm -> Context.generic -> Context.generic
24047
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    12
  val add: attribute
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    13
  val del: attribute
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    14
  val setup: theory -> theory
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    15
end;
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    16
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    17
functor NamedThmsFun(val name: string val description: string): NAMED_THMS =
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    18
struct
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    19
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    20
structure Data = GenericDataFun
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    21
(
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    22
  type T = thm list;
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    23
  val empty = [];
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    24
  val extend = I;
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    25
  fun merge _ = Thm.merge_thms;
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    26
);
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    27
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    28
val get = Data.get o Context.Proof;
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    29
26363
9d95309f8069 export add/del_thm;
wenzelm
parents: 24867
diff changeset
    30
val add_thm = Data.map o Thm.add_thm;
9d95309f8069 export add/del_thm;
wenzelm
parents: 24867
diff changeset
    31
val del_thm = Data.map o Thm.del_thm;
9d95309f8069 export add/del_thm;
wenzelm
parents: 24867
diff changeset
    32
9d95309f8069 export add/del_thm;
wenzelm
parents: 24867
diff changeset
    33
val add = Thm.declaration_attribute add_thm;
9d95309f8069 export add/del_thm;
wenzelm
parents: 24867
diff changeset
    34
val del = Thm.declaration_attribute del_thm;
24047
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    35
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    36
val setup =
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 29579
diff changeset
    37
  Attrib.setup (Binding.name name) (Attrib.add_del add del) ("declaration of " ^ description) #>
29579
cb520b766e00 binding replaces bstring
haftmann
parents: 26724
diff changeset
    38
  PureThy.add_thms_dynamic (Binding.name name, Data.get);
24047
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    39
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    40
end;