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