author | wenzelm |
Mon, 27 Feb 2023 15:09:59 +0100 | |
changeset 77398 | 19e9cafaafc5 |
parent 74563 | 042041c0ebeb |
child 77904 | e7fd273657f1 |
permissions | -rw-r--r-- |
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
1 |
(* Title: Pure/Tools/named_theorems.ML |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
2 |
Author: Makarius |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
3 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
4 |
Named collections of theorems in canonical order. |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
5 |
*) |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
6 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
7 |
signature NAMED_THEOREMS = |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
8 |
sig |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
9 |
val member: Proof.context -> string -> thm -> bool |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
10 |
val get: Proof.context -> string -> thm list |
61900 | 11 |
val clear: string -> Context.generic -> Context.generic |
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
12 |
val add_thm: string -> thm -> Context.generic -> Context.generic |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
13 |
val del_thm: string -> thm -> Context.generic -> Context.generic |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
14 |
val add: string -> attribute |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
15 |
val del: string -> attribute |
59920
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
16 |
val check: Proof.context -> string * Position.T -> string |
70182
ca9dfa7ee3bd
backed out experimental b67bab2b132c, which slipped in accidentally
haftmann
parents:
70177
diff
changeset
|
17 |
val declare: binding -> string -> local_theory -> string * local_theory |
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
18 |
end; |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
19 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
20 |
structure Named_Theorems: NAMED_THEOREMS = |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
21 |
struct |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
22 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
23 |
(* context data *) |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
24 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
25 |
structure Data = Generic_Data |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
26 |
( |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
27 |
type T = thm Item_Net.T Symtab.table; |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
28 |
val empty: T = Symtab.empty; |
57965 | 29 |
val merge : T * T -> T = Symtab.join (K Item_Net.merge); |
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
30 |
); |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
31 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
32 |
fun new_entry name = |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
33 |
Data.map (fn data => |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
34 |
if Symtab.defined data name |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
35 |
then error ("Duplicate declaration of named theorems: " ^ quote name) |
74152 | 36 |
else Symtab.update (name, Thm.item_net) data); |
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
37 |
|
59920
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
38 |
fun undeclared name = "Undeclared named theorems " ^ quote name; |
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
39 |
|
69185 | 40 |
val defined_entry = Symtab.defined o Data.get; |
41 |
||
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
42 |
fun the_entry context name = |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
43 |
(case Symtab.lookup (Data.get context) name of |
59920
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
44 |
NONE => error (undeclared name) |
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
45 |
| SOME entry => entry); |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
46 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
47 |
fun map_entry name f context = |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
48 |
(the_entry context name; Data.map (Symtab.map_entry name f) context); |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
49 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
50 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
51 |
(* maintain content *) |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
52 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
53 |
fun member ctxt = Item_Net.member o the_entry (Context.Proof ctxt); |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
54 |
|
61049 | 55 |
fun content context = |
67649 | 56 |
rev o map (Thm.transfer'' context) o Item_Net.content o the_entry context; |
61049 | 57 |
|
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
58 |
val get = content o Context.Proof; |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
59 |
|
74152 | 60 |
fun clear name = map_entry name (K Thm.item_net); |
61900 | 61 |
|
61049 | 62 |
fun add_thm name th = map_entry name (Item_Net.update (Thm.trim_context th)); |
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
63 |
fun del_thm name = map_entry name o Item_Net.remove; |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
64 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
65 |
val add = Thm.declaration_attribute o add_thm; |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
66 |
val del = Thm.declaration_attribute o del_thm; |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
67 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
68 |
|
59920
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
69 |
(* check *) |
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
70 |
|
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
71 |
fun check ctxt (xname, pos) = |
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
72 |
let |
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
73 |
val context = Context.Proof ctxt; |
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
74 |
val fact_ref = Facts.Named ((xname, Position.none), NONE); |
69185 | 75 |
fun err () = |
76 |
let |
|
77 |
val space = Facts.space_of (Proof_Context.facts_of ctxt); |
|
78 |
val completion = Name_Space.completion context space (defined_entry context) (xname, pos); |
|
69289 | 79 |
in error (undeclared xname ^ Position.here pos ^ Completion.markup_report [completion]) end; |
59920
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
80 |
in |
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
81 |
(case try (Proof_Context.get_fact_generic context) fact_ref of |
69185 | 82 |
SOME (SOME name, _) => if defined_entry context name then name else err () |
59920
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
83 |
| _ => err ()) |
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
84 |
end; |
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
85 |
|
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
86 |
|
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
87 |
(* declaration *) |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
88 |
|
70182
ca9dfa7ee3bd
backed out experimental b67bab2b132c, which slipped in accidentally
haftmann
parents:
70177
diff
changeset
|
89 |
fun declare binding descr lthy = |
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
90 |
let |
59879 | 91 |
val name = Local_Theory.full_name lthy binding; |
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
92 |
val description = |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
93 |
"declaration of " ^ (if descr = "" then Binding.name_of binding ^ " rules" else descr); |
57888 | 94 |
val lthy' = lthy |
57929
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57928
diff
changeset
|
95 |
|> Local_Theory.background_theory (Context.theory_map (new_entry name)) |
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57928
diff
changeset
|
96 |
|> Local_Theory.map_contexts (K (Context.proof_map (new_entry name))) |
70182
ca9dfa7ee3bd
backed out experimental b67bab2b132c, which slipped in accidentally
haftmann
parents:
70177
diff
changeset
|
97 |
|> Local_Theory.add_thms_dynamic (binding, fn context => content context name) |
57927 | 98 |
|> Attrib.local_setup binding (Attrib.add_del (add name) (del name)) description |
57888 | 99 |
in (name, lthy') end; |
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
100 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
101 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
102 |
(* ML antiquotation *) |
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
103 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
104 |
val _ = Theory.setup |
69592
a80d8ec6c998
support for isabelle update -u control_cartouches;
wenzelm
parents:
69349
diff
changeset
|
105 |
(ML_Antiquotation.inline_embedded \<^binding>\<open>named_theorems\<close> |
74563 | 106 |
(Args.context -- Scan.lift Parse.embedded_position >> |
59920
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
wenzelm
parents:
59879
diff
changeset
|
107 |
(fn (ctxt, name) => ML_Syntax.print_string (check ctxt name)))); |
57886
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
108 |
|
7cae177c9084
support for named collections of theorems in canonical order;
wenzelm
parents:
diff
changeset
|
109 |
end; |