src/Pure/Tools/named_thms.ML
author Fabian Huch <huch@in.tum.de>
Thu, 09 Nov 2023 11:41:19 +0100
changeset 78929 df323f23dfde
parent 77905 acee6c7fafff
permissions -rw-r--r--
performance tuning for timing heuristic: pre-calculate graph operations;
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
77905
acee6c7fafff proper Thm.trim_context / Thm.transfer (see also 0d401f874942);
wenzelm
parents: 74561
diff changeset
     4
Named collections of theorems in canonical (reverse) order: OLD VERSION.
24047
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
36296
5cc547abd995 Item_Net/Named_Thms: export efficient member operation;
wenzelm
parents: 33519
diff changeset
     9
  val member: Proof.context -> thm -> bool
24047
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    10
  val get: Proof.context -> thm list
26363
9d95309f8069 export add/del_thm;
wenzelm
parents: 24867
diff changeset
    11
  val add_thm: thm -> Context.generic -> Context.generic
9d95309f8069 export add/del_thm;
wenzelm
parents: 24867
diff changeset
    12
  val del_thm: thm -> Context.generic -> Context.generic
24047
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    13
  val add: attribute
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    14
  val del: attribute
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    15
  val setup: theory -> theory
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    16
end;
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    17
45294
3c5d3d286055 tuned Named_Thms: proper binding;
wenzelm
parents: 39557
diff changeset
    18
functor Named_Thms(val name: binding val description: string): NAMED_THMS =
24047
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    19
struct
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    20
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33453
diff changeset
    21
structure Data = Generic_Data
24047
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    22
(
33453
fe551dc9d4bd scalable version of Named_Thms, using Item_Net;
wenzelm
parents: 33307
diff changeset
    23
  type T = thm Item_Net.T;
74152
069f6b2c5a07 tuned signature;
wenzelm
parents: 45294
diff changeset
    24
  val empty = Thm.item_net;
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33453
diff changeset
    25
  val merge = Item_Net.merge;
24047
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
36296
5cc547abd995 Item_Net/Named_Thms: export efficient member operation;
wenzelm
parents: 33519
diff changeset
    28
val member = Item_Net.member o Data.get o Context.Proof;
5cc547abd995 Item_Net/Named_Thms: export efficient member operation;
wenzelm
parents: 33519
diff changeset
    29
77905
acee6c7fafff proper Thm.trim_context / Thm.transfer (see also 0d401f874942);
wenzelm
parents: 74561
diff changeset
    30
fun content context = map (Thm.transfer'' context) (Item_Net.content (Data.get context));
33453
fe551dc9d4bd scalable version of Named_Thms, using Item_Net;
wenzelm
parents: 33307
diff changeset
    31
val get = content o Context.Proof;
24047
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    32
77905
acee6c7fafff proper Thm.trim_context / Thm.transfer (see also 0d401f874942);
wenzelm
parents: 74561
diff changeset
    33
val add_thm = Data.map o Item_Net.update o Thm.trim_context;
33453
fe551dc9d4bd scalable version of Named_Thms, using Item_Net;
wenzelm
parents: 33307
diff changeset
    34
val del_thm = Data.map o Item_Net.remove;
26363
9d95309f8069 export add/del_thm;
wenzelm
parents: 24867
diff changeset
    35
9d95309f8069 export add/del_thm;
wenzelm
parents: 24867
diff changeset
    36
val add = Thm.declaration_attribute add_thm;
9d95309f8069 export add/del_thm;
wenzelm
parents: 24867
diff changeset
    37
val del = Thm.declaration_attribute del_thm;
24047
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    38
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    39
val setup =
45294
3c5d3d286055 tuned Named_Thms: proper binding;
wenzelm
parents: 39557
diff changeset
    40
  Attrib.setup name (Attrib.add_del add del) ("declaration of " ^ description) #>
3c5d3d286055 tuned Named_Thms: proper binding;
wenzelm
parents: 39557
diff changeset
    41
  Global_Theory.add_thms_dynamic (name, content);
24047
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    42
47b588ce11ec Named collections of theorems in canonical order.
wenzelm
parents:
diff changeset
    43
end;