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