| author | wenzelm | 
| Wed, 13 Jan 2016 23:25:18 +0100 | |
| changeset 62176 | 1221f2a46945 | 
| parent 61900 | 3f5e2e0a6d29 | 
| child 62848 | e4140efe699e | 
| 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  | 
| 57888 | 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;  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
29  | 
val extend = I;  | 
| 57965 | 30  | 
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
 | 
31  | 
);  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
32  | 
|
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
33  | 
fun new_entry name =  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
34  | 
Data.map (fn data =>  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
if Symtab.defined data name  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
36  | 
    then error ("Duplicate declaration of named theorems: " ^ quote name)
 | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
else Symtab.update (name, Thm.full_rules) data);  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
|
| 
59920
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
39  | 
fun undeclared name = "Undeclared named theorems " ^ quote name;  | 
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
40  | 
|
| 
57886
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
41  | 
fun the_entry context name =  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
42  | 
(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
 | 
43  | 
NONE => error (undeclared name)  | 
| 
57886
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
44  | 
| SOME entry => entry);  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
|
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
fun map_entry name f context =  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
(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
 | 
48  | 
|
| 
 
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  | 
(* maintain content *)  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
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
 | 
53  | 
|
| 61049 | 54  | 
fun content context =  | 
55  | 
rev o map (Thm.transfer (Context.theory_of context)) o Item_Net.content o the_entry context;  | 
|
56  | 
||
| 
57886
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
57  | 
val get = content o Context.Proof;  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
58  | 
|
| 61900 | 59  | 
fun clear name = map_entry name (K Thm.full_rules);  | 
60  | 
||
| 61049 | 61  | 
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
 | 
62  | 
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
 | 
63  | 
|
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
64  | 
val add = Thm.declaration_attribute o add_thm;  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
65  | 
val del = Thm.declaration_attribute o del_thm;  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
66  | 
|
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
67  | 
|
| 
59920
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
68  | 
(* check *)  | 
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
69  | 
|
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
70  | 
fun check ctxt (xname, pos) =  | 
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
71  | 
let  | 
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
72  | 
val context = Context.Proof ctxt;  | 
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
73  | 
val fact_ref = Facts.Named ((xname, Position.none), NONE);  | 
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
74  | 
fun err () = error (undeclared xname ^ Position.here pos);  | 
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
75  | 
in  | 
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
76  | 
(case try (Proof_Context.get_fact_generic context) fact_ref of  | 
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
77  | 
SOME (SOME name, _) => if can (the_entry context) name then name else err ()  | 
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
78  | 
| _ => err ())  | 
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
79  | 
end;  | 
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
80  | 
|
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
81  | 
|
| 
57886
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
82  | 
(* declaration *)  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
83  | 
|
| 57888 | 84  | 
fun declare binding descr lthy =  | 
| 
57886
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
85  | 
let  | 
| 59879 | 86  | 
val name = Local_Theory.full_name lthy binding;  | 
| 
57886
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
87  | 
val description =  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
88  | 
"declaration of " ^ (if descr = "" then Binding.name_of binding ^ " rules" else descr);  | 
| 57888 | 89  | 
val lthy' = lthy  | 
| 
57929
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
90  | 
|> Local_Theory.background_theory (Context.theory_map (new_entry name))  | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
91  | 
|> Local_Theory.map_contexts (K (Context.proof_map (new_entry name)))  | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
92  | 
|> Local_Theory.add_thms_dynamic (binding, fn context => content context name)  | 
| 57927 | 93  | 
|> Attrib.local_setup binding (Attrib.add_del (add name) (del name)) description  | 
| 57888 | 94  | 
in (name, lthy') end;  | 
| 
57886
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
95  | 
|
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
96  | 
val _ =  | 
| 
59936
 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 
wenzelm 
parents: 
59920 
diff
changeset
 | 
97  | 
  Outer_Syntax.local_theory @{command_keyword named_theorems}
 | 
| 
57886
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
98  | 
"declare named collection of theorems"  | 
| 59028 | 99  | 
(Parse.and_list1 (Parse.binding -- Scan.optional Parse.text "") >>  | 
100  | 
fold (fn (b, descr) => snd o declare b descr));  | 
|
| 
57886
 
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  | 
|
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
103  | 
(* ML antiquotation *)  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
104  | 
|
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
105  | 
val _ = Theory.setup  | 
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
106  | 
  (ML_Antiquotation.inline @{binding named_theorems}
 | 
| 
59920
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
107  | 
(Args.context -- Scan.lift (Parse.position Args.name) >>  | 
| 
 
86d302846b16
check wrt. proper context, e.g. relevant for 'experiment' target;
 
wenzelm 
parents: 
59879 
diff
changeset
 | 
108  | 
(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
 | 
109  | 
|
| 
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents:  
diff
changeset
 | 
110  | 
end;  |