| author | wenzelm | 
| Mon, 16 Mar 2009 18:24:39 +0100 | |
| changeset 30550 | c601204b055c | 
| parent 30543 | 2ca69af4422a | 
| child 30575 | 368e26dfba69 | 
| permissions | -rw-r--r-- | 
| 5823 | 1  | 
(* Title: Pure/Isar/attrib.ML  | 
2  | 
Author: Markus Wenzel, TU Muenchen  | 
|
3  | 
||
| 18734 | 4  | 
Symbolic representation of attributes -- with name and syntax.  | 
| 5823 | 5  | 
*)  | 
6  | 
||
7  | 
signature ATTRIB =  | 
|
8  | 
sig  | 
|
| 27729 | 9  | 
type src = Args.src  | 
| 29581 | 10  | 
type binding = binding * src list  | 
| 28965 | 11  | 
val empty_binding: binding  | 
| 27729 | 12  | 
val print_attributes: theory -> unit  | 
| 16458 | 13  | 
val intern: theory -> xstring -> string  | 
14  | 
val intern_src: theory -> src -> src  | 
|
| 21031 | 15  | 
val pretty_attribs: Proof.context -> src list -> Pretty.T list  | 
| 26891 | 16  | 
val defined: theory -> string -> bool  | 
| 18734 | 17  | 
val attribute: theory -> src -> attribute  | 
18  | 
val attribute_i: theory -> src -> attribute  | 
|
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
19  | 
val eval_thms: Proof.context -> (Facts.ref * src list) list -> thm list  | 
| 18905 | 20  | 
  val map_specs: ('a -> 'att) ->
 | 
21  | 
    (('c * 'a list) * 'd) list -> (('c * 'att list) * 'd) list
 | 
|
22  | 
  val map_facts: ('a -> 'att) ->
 | 
|
| 
17105
 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 
wenzelm 
parents: 
16934 
diff
changeset
 | 
23  | 
    (('c * 'a list) * ('d * 'a list) list) list ->
 | 
| 18905 | 24  | 
    (('c * 'att list) * ('d * 'att list) list) list
 | 
| 20289 | 25  | 
val crude_closure: Proof.context -> src -> src  | 
| 18734 | 26  | 
val add_attributes: (bstring * (src -> attribute) * string) list -> theory -> theory  | 
| 30513 | 27  | 
val syntax: attribute context_parser -> src -> attribute  | 
| 
30525
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
28  | 
val setup: Binding.binding -> attribute context_parser -> string -> theory -> theory  | 
| 30543 | 29  | 
val attribute_setup: bstring * Position.T -> string * Position.T -> string -> theory -> theory  | 
| 25983 | 30  | 
val no_args: attribute -> src -> attribute  | 
| 
30525
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
31  | 
val add_del: attribute -> attribute -> attribute context_parser  | 
| 25983 | 32  | 
val add_del_args: attribute -> attribute -> src -> attribute  | 
| 30513 | 33  | 
val thm_sel: Facts.interval list parser  | 
34  | 
val thm: thm context_parser  | 
|
35  | 
val thms: thm list context_parser  | 
|
36  | 
val multi_thm: thm list context_parser  | 
|
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
37  | 
val print_configs: Proof.context -> unit  | 
| 25983 | 38  | 
val internal: (morphism -> attribute) -> src  | 
| 
24126
 
913e1fa904fb
turned simp_depth_limit into configuration option;
 
wenzelm 
parents: 
24110 
diff
changeset
 | 
39  | 
val register_config: Config.value Config.T -> theory -> theory  | 
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
40  | 
val config_bool: bstring -> bool -> bool Config.T * (theory -> theory)  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
41  | 
val config_int: bstring -> int -> int Config.T * (theory -> theory)  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
42  | 
val config_string: bstring -> string -> string Config.T * (theory -> theory)  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
43  | 
val config_bool_global: bstring -> bool -> bool Config.T * (theory -> theory)  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
44  | 
val config_int_global: bstring -> int -> int Config.T * (theory -> theory)  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
45  | 
val config_string_global: bstring -> string -> string Config.T * (theory -> theory)  | 
| 5823 | 46  | 
end;  | 
47  | 
||
48  | 
structure Attrib: ATTRIB =  | 
|
49  | 
struct  | 
|
50  | 
||
| 
27812
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
51  | 
structure T = OuterLex;  | 
| 
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
52  | 
structure P = OuterParse;  | 
| 
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
53  | 
|
| 
28084
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28083 
diff
changeset
 | 
54  | 
|
| 
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28083 
diff
changeset
 | 
55  | 
(* source and bindings *)  | 
| 
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28083 
diff
changeset
 | 
56  | 
|
| 15703 | 57  | 
type src = Args.src;  | 
58  | 
||
| 29581 | 59  | 
type binding = binding * src list;  | 
| 28965 | 60  | 
val empty_binding: binding = (Binding.empty, []);  | 
| 
28084
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28083 
diff
changeset
 | 
61  | 
|
| 
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28083 
diff
changeset
 | 
62  | 
|
| 27729 | 63  | 
|
| 18734 | 64  | 
(** named attributes **)  | 
| 18636 | 65  | 
|
| 18734 | 66  | 
(* theory data *)  | 
| 5823 | 67  | 
|
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
68  | 
structure Attributes = TheoryDataFun  | 
| 22846 | 69  | 
(  | 
| 18734 | 70  | 
type T = (((src -> attribute) * string) * stamp) NameSpace.table;  | 
| 16344 | 71  | 
val empty = NameSpace.empty_table;  | 
| 6546 | 72  | 
val copy = I;  | 
| 16458 | 73  | 
val extend = I;  | 
| 
23655
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
23577 
diff
changeset
 | 
74  | 
fun merge _ tables : T = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUP dup =>  | 
| 
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
23577 
diff
changeset
 | 
75  | 
    error ("Attempt to merge different versions of attribute " ^ quote dup);
 | 
| 22846 | 76  | 
);  | 
| 5823 | 77  | 
|
| 22846 | 78  | 
fun print_attributes thy =  | 
79  | 
let  | 
|
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
80  | 
val attribs = Attributes.get thy;  | 
| 22846 | 81  | 
fun prt_attr (name, ((_, comment), _)) = Pretty.block  | 
82  | 
[Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];  | 
|
83  | 
in  | 
|
84  | 
[Pretty.big_list "attributes:" (map prt_attr (NameSpace.extern_table attribs))]  | 
|
85  | 
|> Pretty.chunks |> Pretty.writeln  | 
|
86  | 
end;  | 
|
| 7611 | 87  | 
|
| 5823 | 88  | 
|
| 21031 | 89  | 
(* name space *)  | 
| 15703 | 90  | 
|
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
91  | 
val intern = NameSpace.intern o #1 o Attributes.get;  | 
| 15703 | 92  | 
val intern_src = Args.map_name o intern;  | 
93  | 
||
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
94  | 
val extern = NameSpace.extern o #1 o Attributes.get o ProofContext.theory_of;  | 
| 21031 | 95  | 
|
96  | 
||
97  | 
(* pretty printing *)  | 
|
98  | 
||
99  | 
fun pretty_attribs _ [] = []  | 
|
100  | 
| pretty_attribs ctxt srcs =  | 
|
101  | 
[Pretty.enclose "[" "]"  | 
|
102  | 
(Pretty.commas (map (Args.pretty_src ctxt o Args.map_name (extern ctxt)) srcs))];  | 
|
103  | 
||
| 15703 | 104  | 
|
| 18734 | 105  | 
(* get attributes *)  | 
| 5823 | 106  | 
|
| 26891 | 107  | 
val defined = Symtab.defined o #2 o Attributes.get;  | 
108  | 
||
| 18734 | 109  | 
fun attribute_i thy =  | 
| 5823 | 110  | 
let  | 
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
111  | 
val attrs = #2 (Attributes.get thy);  | 
| 5879 | 112  | 
fun attr src =  | 
| 16344 | 113  | 
let val ((name, _), pos) = Args.dest_src src in  | 
| 17412 | 114  | 
(case Symtab.lookup attrs name of  | 
| 15531 | 115  | 
          NONE => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
 | 
| 27751 | 116  | 
| SOME ((att, _), _) => (Position.report (Markup.attribute name) pos; att src))  | 
| 5823 | 117  | 
end;  | 
118  | 
in attr end;  | 
|
119  | 
||
| 18734 | 120  | 
fun attribute thy = attribute_i thy o intern_src thy;  | 
| 18636 | 121  | 
|
| 24723 | 122  | 
fun eval_thms ctxt args = ProofContext.note_thmss Thm.theoremK  | 
| 30211 | 123  | 
[(Thm.empty_binding, map (apsnd (map (attribute (ProofContext.theory_of ctxt)))) args)] ctxt  | 
| 26685 | 124  | 
|> fst |> maps snd;  | 
| 24723 | 125  | 
|
| 5823 | 126  | 
|
| 
17105
 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 
wenzelm 
parents: 
16934 
diff
changeset
 | 
127  | 
(* attributed declarations *)  | 
| 
 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 
wenzelm 
parents: 
16934 
diff
changeset
 | 
128  | 
|
| 
 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 
wenzelm 
parents: 
16934 
diff
changeset
 | 
129  | 
fun map_specs f = map (apfst (apsnd (map f)));  | 
| 
 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 
wenzelm 
parents: 
16934 
diff
changeset
 | 
130  | 
fun map_facts f = map (apfst (apsnd (map f)) o apsnd (map (apsnd (map f))));  | 
| 
 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 
wenzelm 
parents: 
16934 
diff
changeset
 | 
131  | 
|
| 
 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 
wenzelm 
parents: 
16934 
diff
changeset
 | 
132  | 
|
| 15703 | 133  | 
(* crude_closure *)  | 
134  | 
||
135  | 
(*Produce closure without knowing facts in advance! The following  | 
|
| 18734 | 136  | 
works reasonably well for attribute parsers that do not peek at the  | 
137  | 
thm structure.*)  | 
|
| 15703 | 138  | 
|
139  | 
fun crude_closure ctxt src =  | 
|
| 18734 | 140  | 
(try (fn () => attribute_i (ProofContext.theory_of ctxt) src  | 
141  | 
(Context.Proof ctxt, Drule.asm_rl)) ();  | 
|
| 15703 | 142  | 
Args.closure src);  | 
143  | 
||
144  | 
||
| 5823 | 145  | 
(* add_attributes *)  | 
146  | 
||
147  | 
fun add_attributes raw_attrs thy =  | 
|
148  | 
let  | 
|
| 18734 | 149  | 
val new_attrs =  | 
| 29004 | 150  | 
raw_attrs |> map (fn (name, att, comment) => (Binding.name name, ((att, comment), stamp ())));  | 
| 30466 | 151  | 
fun add attrs = fold (snd oo NameSpace.define (Sign.naming_of thy)) new_attrs attrs  | 
| 
23655
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
23577 
diff
changeset
 | 
152  | 
      handle Symtab.DUP dup => error ("Duplicate declaration of attributes " ^ quote dup);
 | 
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
153  | 
in Attributes.map add thy end;  | 
| 5823 | 154  | 
|
| 5879 | 155  | 
|
| 
30525
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
156  | 
(* attribute setup *)  | 
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
157  | 
|
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
158  | 
fun syntax scan src (context, th) =  | 
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
159  | 
let val (f: attribute, context') = Args.syntax "attribute" scan src context  | 
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
160  | 
in f (context', th) end;  | 
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
161  | 
|
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
162  | 
fun setup name scan comment = add_attributes [(Binding.name_of name, syntax scan, comment)];  | 
| 5879 | 163  | 
|
| 
30525
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
164  | 
fun attribute_setup name (txt, pos) cmt =  | 
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
165  | 
Context.theory_map (ML_Context.expression pos  | 
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
166  | 
"val (name, scan, comment): binding * attribute context_parser * string"  | 
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
167  | 
"Context.map_theory (Attrib.setup name scan comment)"  | 
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
168  | 
    ("(" ^ ML_Syntax.make_binding name ^ ", " ^ txt ^ ", " ^ ML_Syntax.print_string cmt ^ ")"));
 | 
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
169  | 
|
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
170  | 
|
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
171  | 
(* basic syntax *)  | 
| 5823 | 172  | 
|
| 25983 | 173  | 
fun no_args x = syntax (Scan.succeed x);  | 
174  | 
||
| 
30525
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
175  | 
fun add_del add del = (Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add));  | 
| 
 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 
wenzelm 
parents: 
30513 
diff
changeset
 | 
176  | 
fun add_del_args add del = syntax (add_del add del);  | 
| 5879 | 177  | 
|
178  | 
||
| 25983 | 179  | 
|
180  | 
(** parsing attributed theorems **)  | 
|
| 5879 | 181  | 
|
| 
27812
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
182  | 
val thm_sel = P.$$$ "(" |-- P.list1
 | 
| 
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
183  | 
(P.nat --| P.minus -- P.nat >> Facts.FromTo ||  | 
| 
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
184  | 
P.nat --| P.minus >> Facts.From ||  | 
| 
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
185  | 
P.nat >> Facts.Single) --| P.$$$ ")";  | 
| 
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
186  | 
|
| 18636 | 187  | 
local  | 
188  | 
||
| 
21698
 
43a842769765
thms etc.: proper treatment of internal_fact with selection;
 
wenzelm 
parents: 
21658 
diff
changeset
 | 
189  | 
val fact_name = Args.internal_fact >> K "<fact>" || Args.name;  | 
| 
 
43a842769765
thms etc.: proper treatment of internal_fact with selection;
 
wenzelm 
parents: 
21658 
diff
changeset
 | 
190  | 
|
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
191  | 
fun gen_thm pick = Scan.depend (fn context =>  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
192  | 
let  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
193  | 
val thy = Context.theory_of context;  | 
| 26685 | 194  | 
val get = Context.cases (PureThy.get_fact context) ProofContext.get_fact context;  | 
| 26361 | 195  | 
val get_fact = get o Facts.Fact;  | 
| 27820 | 196  | 
fun get_named pos name = get (Facts.Named ((name, pos), NONE));  | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
197  | 
in  | 
| 27820 | 198  | 
P.$$$ "[" |-- Args.attribs (intern thy) --| P.$$$ "]" >> (fn srcs =>  | 
| 24008 | 199  | 
let  | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
200  | 
val atts = map (attribute_i thy) srcs;  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
201  | 
val (context', th') = Library.apply atts (context, Drule.dummy_thm);  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
202  | 
in (context', pick "" [th']) end)  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
203  | 
||  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
204  | 
(Scan.ahead Args.alt_name -- Args.named_fact get_fact  | 
| 27820 | 205  | 
      >> (fn (s, fact) => ("", Facts.Fact s, fact)) ||
 | 
206  | 
Scan.ahead (P.position fact_name) :|-- (fn (name, pos) =>  | 
|
207  | 
Args.named_fact (get_named pos) -- Scan.option thm_sel  | 
|
208  | 
>> (fn (fact, sel) => (name, Facts.Named ((name, pos), sel), fact))))  | 
|
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
209  | 
-- Args.opt_attribs (intern thy) >> (fn ((name, thmref, fact), srcs) =>  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
210  | 
let  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
211  | 
val ths = Facts.select thmref fact;  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
212  | 
val atts = map (attribute_i thy) srcs;  | 
| 30190 | 213  | 
val (context', ths') = Library.foldl_map (Library.apply atts) (context, ths);  | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
214  | 
in (context', pick name ths') end)  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
215  | 
end);  | 
| 
15456
 
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
 
berghofe 
parents: 
15117 
diff
changeset
 | 
216  | 
|
| 18636 | 217  | 
in  | 
218  | 
||
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25983 
diff
changeset
 | 
219  | 
val thm = gen_thm Facts.the_single;  | 
| 18998 | 220  | 
val multi_thm = gen_thm (K I);  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19473 
diff
changeset
 | 
221  | 
val thms = Scan.repeat multi_thm >> flat;  | 
| 18636 | 222  | 
|
223  | 
end;  | 
|
224  | 
||
| 5823 | 225  | 
|
| 5879 | 226  | 
|
| 25983 | 227  | 
(** basic attributes **)  | 
228  | 
||
229  | 
(* internal *)  | 
|
230  | 
||
| 
27812
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
231  | 
fun internal att = Args.src (("Pure.attribute", [T.mk_attribute att]), Position.none);
 | 
| 25983 | 232  | 
|
233  | 
val internal_att =  | 
|
234  | 
syntax (Scan.lift Args.internal_attribute >> Morphism.form);  | 
|
235  | 
||
236  | 
||
237  | 
(* tags *)  | 
|
238  | 
||
| 
27865
 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 
wenzelm 
parents: 
27820 
diff
changeset
 | 
239  | 
val tagged = syntax (Scan.lift (Args.name -- Args.name) >> Thm.tag);  | 
| 
 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 
wenzelm 
parents: 
27820 
diff
changeset
 | 
240  | 
val untagged = syntax (Scan.lift Args.name >> Thm.untag);  | 
| 25983 | 241  | 
|
| 
27865
 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 
wenzelm 
parents: 
27820 
diff
changeset
 | 
242  | 
val kind = syntax (Scan.lift Args.name >> Thm.kind);  | 
| 25983 | 243  | 
|
244  | 
||
245  | 
(* rule composition *)  | 
|
246  | 
||
247  | 
val COMP_att =  | 
|
| 
27812
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
248  | 
syntax (Scan.lift (Scan.optional (Args.bracks P.nat) 1) -- thm  | 
| 25983 | 249  | 
>> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => Drule.compose_single (A, i, B))));  | 
250  | 
||
251  | 
val THEN_att =  | 
|
| 
27812
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
252  | 
syntax (Scan.lift (Scan.optional (Args.bracks P.nat) 1) -- thm  | 
| 25983 | 253  | 
>> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => A RSN (i, B))));  | 
254  | 
||
255  | 
val OF_att =  | 
|
256  | 
syntax (thms >> (fn Bs => Thm.rule_attribute (fn _ => fn A => Bs MRS A)));  | 
|
257  | 
||
258  | 
||
259  | 
(* rename_abs *)  | 
|
260  | 
||
261  | 
val rename_abs = syntax  | 
|
262  | 
(Scan.lift (Scan.repeat (Args.maybe Args.name) >> (apsnd o Drule.rename_bvars')));  | 
|
263  | 
||
264  | 
||
265  | 
(* unfold / fold definitions *)  | 
|
266  | 
||
267  | 
fun unfolded_syntax rule =  | 
|
268  | 
syntax (thms >>  | 
|
269  | 
(fn ths => Thm.rule_attribute (fn context => rule (Context.proof_of context) ths)));  | 
|
270  | 
||
271  | 
val unfolded = unfolded_syntax LocalDefs.unfold;  | 
|
272  | 
val folded = unfolded_syntax LocalDefs.fold;  | 
|
273  | 
||
274  | 
||
275  | 
(* rule cases *)  | 
|
| 5823 | 276  | 
|
| 
27812
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
277  | 
val consumes = syntax (Scan.lift (Scan.optional P.nat 1) >> RuleCases.consumes);  | 
| 25983 | 278  | 
val case_names = syntax (Scan.lift (Scan.repeat1 Args.name) >> RuleCases.case_names);  | 
279  | 
val case_conclusion =  | 
|
280  | 
syntax (Scan.lift (Args.name -- Scan.repeat Args.name) >> RuleCases.case_conclusion);  | 
|
| 
27812
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
281  | 
val params = syntax (Scan.lift (P.and_list1 (Scan.repeat Args.name)) >> RuleCases.params);  | 
| 25983 | 282  | 
|
283  | 
||
284  | 
(* rule format *)  | 
|
285  | 
||
286  | 
val rule_format = syntax (Args.mode "no_asm"  | 
|
287  | 
>> (fn true => ObjectLogic.rule_format_no_asm | false => ObjectLogic.rule_format));  | 
|
288  | 
||
289  | 
val elim_format = no_args (Thm.rule_attribute (K Tactic.make_elim));  | 
|
290  | 
||
291  | 
||
292  | 
(* misc rules *)  | 
|
293  | 
||
294  | 
val standard = no_args (Thm.rule_attribute (K Drule.standard));  | 
|
295  | 
||
| 
26715
 
00ff79ab6e6b
no_vars: reset body mode, i.e. invent global frees (which are acceptable to Variable.auto_fixes);
 
wenzelm 
parents: 
26685 
diff
changeset
 | 
296  | 
val no_vars = no_args (Thm.rule_attribute (fn context => fn th =>  | 
| 
 
00ff79ab6e6b
no_vars: reset body mode, i.e. invent global frees (which are acceptable to Variable.auto_fixes);
 
wenzelm 
parents: 
26685 
diff
changeset
 | 
297  | 
let  | 
| 
 
00ff79ab6e6b
no_vars: reset body mode, i.e. invent global frees (which are acceptable to Variable.auto_fixes);
 
wenzelm 
parents: 
26685 
diff
changeset
 | 
298  | 
val ctxt = Variable.set_body false (Context.proof_of context);  | 
| 
 
00ff79ab6e6b
no_vars: reset body mode, i.e. invent global frees (which are acceptable to Variable.auto_fixes);
 
wenzelm 
parents: 
26685 
diff
changeset
 | 
299  | 
val ((_, [th']), _) = Variable.import_thms true [th] ctxt;  | 
| 25983 | 300  | 
in th' end));  | 
301  | 
||
302  | 
val eta_long =  | 
|
303  | 
no_args (Thm.rule_attribute (K (Conv.fconv_rule Drule.eta_long_conversion)));  | 
|
304  | 
||
305  | 
val rotated = syntax  | 
|
| 
27812
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
306  | 
(Scan.lift (Scan.optional P.int 1) >> (fn n => Thm.rule_attribute (K (rotate_prems n))));  | 
| 5879 | 307  | 
|
| 29690 | 308  | 
val abs_def = no_args (Thm.rule_attribute (K Drule.abs_def));  | 
309  | 
||
| 25983 | 310  | 
|
311  | 
(* theory setup *)  | 
|
| 5823 | 312  | 
|
| 26463 | 313  | 
val _ = Context.>> (Context.map_theory  | 
| 25983 | 314  | 
(add_attributes  | 
315  | 
   [("attribute", internal_att, "internal attribute"),
 | 
|
316  | 
    ("tagged", tagged, "tagged theorem"),
 | 
|
317  | 
    ("untagged", untagged, "untagged theorem"),
 | 
|
318  | 
    ("kind", kind, "theorem kind"),
 | 
|
319  | 
    ("COMP", COMP_att, "direct composition with rules (no lifting)"),
 | 
|
320  | 
    ("THEN", THEN_att, "resolution with rule"),
 | 
|
321  | 
    ("OF", OF_att, "rule applied to facts"),
 | 
|
322  | 
    ("rename_abs", rename_abs, "rename bound variables in abstractions"),
 | 
|
323  | 
    ("unfolded", unfolded, "unfolded definitions"),
 | 
|
324  | 
    ("folded", folded, "folded definitions"),
 | 
|
325  | 
    ("standard", standard, "result put into standard form"),
 | 
|
326  | 
    ("elim_format", elim_format, "destruct rule turned into elimination rule format"),
 | 
|
327  | 
    ("no_vars", no_vars, "frozen schematic vars"),
 | 
|
328  | 
    ("eta_long", eta_long, "put theorem into eta long beta normal form"),
 | 
|
329  | 
    ("consumes", consumes, "number of consumed facts"),
 | 
|
330  | 
    ("case_names", case_names, "named rule cases"),
 | 
|
331  | 
    ("case_conclusion", case_conclusion, "named conclusion of rule cases"),
 | 
|
332  | 
    ("params", params, "named rule parameters"),
 | 
|
333  | 
    ("atomize", no_args ObjectLogic.declare_atomize, "declaration of atomize rule"),
 | 
|
334  | 
    ("rulify", no_args ObjectLogic.declare_rulify, "declaration of rulify rule"),
 | 
|
335  | 
    ("rule_format", rule_format, "result put into standard rule format"),
 | 
|
336  | 
    ("rotated", rotated, "rotated theorem premises"),
 | 
|
337  | 
    ("defn", add_del_args LocalDefs.defn_add LocalDefs.defn_del,
 | 
|
| 29690 | 338  | 
"declaration of definitional transformations"),  | 
339  | 
    ("abs_def", abs_def, "abstract over free variables of a definition")]));
 | 
|
| 8633 | 340  | 
|
| 5823 | 341  | 
|
342  | 
||
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
343  | 
(** configuration options **)  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
344  | 
|
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
345  | 
(* naming *)  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
346  | 
|
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
347  | 
structure Configs = TheoryDataFun  | 
| 24713 | 348  | 
(  | 
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
349  | 
type T = Config.value Config.T Symtab.table;  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
350  | 
val empty = Symtab.empty;  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
351  | 
val copy = I;  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
352  | 
val extend = I;  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
353  | 
fun merge _ = Symtab.merge (K true);  | 
| 24713 | 354  | 
);  | 
| 5823 | 355  | 
|
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
356  | 
fun print_configs ctxt =  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
357  | 
let  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
358  | 
val thy = ProofContext.theory_of ctxt;  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
359  | 
fun prt (name, config) =  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
360  | 
let val value = Config.get ctxt config in  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
361  | 
Pretty.block [Pretty.str (name ^ ": " ^ Config.print_type value ^ " ="), Pretty.brk 1,  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
362  | 
Pretty.str (Config.print_value value)]  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
363  | 
end;  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
364  | 
val configs = NameSpace.extern_table (#1 (Attributes.get thy), Configs.get thy);  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
365  | 
in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end;  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
366  | 
|
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
367  | 
|
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
368  | 
(* concrete syntax *)  | 
| 
23988
 
aa46577f4f44
added attribute "option" for setting configuration options;
 
wenzelm 
parents: 
23937 
diff
changeset
 | 
369  | 
|
| 
24003
 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 
wenzelm 
parents: 
23988 
diff
changeset
 | 
370  | 
local  | 
| 
 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 
wenzelm 
parents: 
23988 
diff
changeset
 | 
371  | 
|
| 27820 | 372  | 
val equals = P.$$$ "=";  | 
| 
24003
 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 
wenzelm 
parents: 
23988 
diff
changeset
 | 
373  | 
|
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
374  | 
fun scan_value (Config.Bool _) =  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
375  | 
equals -- Args.$$$ "false" >> K (Config.Bool false) ||  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
376  | 
equals -- Args.$$$ "true" >> K (Config.Bool true) ||  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
377  | 
Scan.succeed (Config.Bool true)  | 
| 
27812
 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27751 
diff
changeset
 | 
378  | 
| scan_value (Config.Int _) = equals |-- P.int >> Config.Int  | 
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
379  | 
| scan_value (Config.String _) = equals |-- Args.name >> Config.String;  | 
| 
24003
 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 
wenzelm 
parents: 
23988 
diff
changeset
 | 
380  | 
|
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
381  | 
fun scan_config thy config =  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
382  | 
let val config_type = Config.get_thy thy config  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
383  | 
in scan_value config_type >> (K o Thm.declaration_attribute o K o Config.put_generic config) end;  | 
| 
24003
 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 
wenzelm 
parents: 
23988 
diff
changeset
 | 
384  | 
|
| 
 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 
wenzelm 
parents: 
23988 
diff
changeset
 | 
385  | 
in  | 
| 
 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 
wenzelm 
parents: 
23988 
diff
changeset
 | 
386  | 
|
| 
24126
 
913e1fa904fb
turned simp_depth_limit into configuration option;
 
wenzelm 
parents: 
24110 
diff
changeset
 | 
387  | 
fun register_config config thy =  | 
| 
 
913e1fa904fb
turned simp_depth_limit into configuration option;
 
wenzelm 
parents: 
24110 
diff
changeset
 | 
388  | 
let  | 
| 
 
913e1fa904fb
turned simp_depth_limit into configuration option;
 
wenzelm 
parents: 
24110 
diff
changeset
 | 
389  | 
val bname = Config.name_of config;  | 
| 28965 | 390  | 
val name = Sign.full_bname thy bname;  | 
| 
24126
 
913e1fa904fb
turned simp_depth_limit into configuration option;
 
wenzelm 
parents: 
24110 
diff
changeset
 | 
391  | 
in  | 
| 
 
913e1fa904fb
turned simp_depth_limit into configuration option;
 
wenzelm 
parents: 
24110 
diff
changeset
 | 
392  | 
thy  | 
| 
 
913e1fa904fb
turned simp_depth_limit into configuration option;
 
wenzelm 
parents: 
24110 
diff
changeset
 | 
393  | 
|> add_attributes [(bname, syntax (Scan.lift (scan_config thy config) >> Morphism.form),  | 
| 
 
913e1fa904fb
turned simp_depth_limit into configuration option;
 
wenzelm 
parents: 
24110 
diff
changeset
 | 
394  | 
"configuration option")]  | 
| 
 
913e1fa904fb
turned simp_depth_limit into configuration option;
 
wenzelm 
parents: 
24110 
diff
changeset
 | 
395  | 
|> Configs.map (Symtab.update (name, config))  | 
| 
 
913e1fa904fb
turned simp_depth_limit into configuration option;
 
wenzelm 
parents: 
24110 
diff
changeset
 | 
396  | 
end;  | 
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
397  | 
|
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
398  | 
fun declare_config make coerce global name default =  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
399  | 
let  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
400  | 
val config_value = Config.declare global name (make default);  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
401  | 
val config = coerce config_value;  | 
| 
24126
 
913e1fa904fb
turned simp_depth_limit into configuration option;
 
wenzelm 
parents: 
24110 
diff
changeset
 | 
402  | 
in (config, register_config config_value) end;  | 
| 
24110
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
403  | 
|
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
404  | 
val config_bool = declare_config Config.Bool Config.bool false;  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
405  | 
val config_int = declare_config Config.Int Config.int false;  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
406  | 
val config_string = declare_config Config.String Config.string false;  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
407  | 
|
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
408  | 
val config_bool_global = declare_config Config.Bool Config.bool true;  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
409  | 
val config_int_global = declare_config Config.Int Config.int true;  | 
| 
 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 
wenzelm 
parents: 
24030 
diff
changeset
 | 
410  | 
val config_string_global = declare_config Config.String Config.string true;  | 
| 
24003
 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 
wenzelm 
parents: 
23988 
diff
changeset
 | 
411  | 
|
| 
 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 
wenzelm 
parents: 
23988 
diff
changeset
 | 
412  | 
end;  | 
| 
23988
 
aa46577f4f44
added attribute "option" for setting configuration options;
 
wenzelm 
parents: 
23937 
diff
changeset
 | 
413  | 
|
| 
 
aa46577f4f44
added attribute "option" for setting configuration options;
 
wenzelm 
parents: 
23937 
diff
changeset
 | 
414  | 
|
| 18636 | 415  | 
(* theory setup *)  | 
| 5823 | 416  | 
|
| 26463 | 417  | 
val _ = Context.>> (Context.map_theory  | 
| 
24178
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
24126 
diff
changeset
 | 
418  | 
(register_config Unify.trace_bound_value #>  | 
| 
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
24126 
diff
changeset
 | 
419  | 
register_config Unify.search_bound_value #>  | 
| 
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
24126 
diff
changeset
 | 
420  | 
register_config Unify.trace_simp_value #>  | 
| 
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
24126 
diff
changeset
 | 
421  | 
register_config Unify.trace_types_value #>  | 
| 26463 | 422  | 
register_config MetaSimplifier.simp_depth_limit_value));  | 
| 5823 | 423  | 
|
424  | 
end;  |