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