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