author | wenzelm |
Thu, 27 Apr 2006 15:06:35 +0200 | |
changeset 19482 | 9f11af8f7ef9 |
parent 19473 | d87a8838afa4 |
child 19798 | 94f12468bbba |
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 |
|
18734 | 11 |
val Attribute: bstring -> (Args.src -> attribute) -> string -> unit |
5823 | 12 |
end; |
13 |
||
14 |
signature ATTRIB = |
|
15 |
sig |
|
16 |
include BASIC_ATTRIB |
|
15703 | 17 |
type src |
5912 | 18 |
exception ATTRIB_FAIL of (string * Position.T) * exn |
16458 | 19 |
val intern: theory -> xstring -> string |
20 |
val intern_src: theory -> src -> src |
|
18734 | 21 |
val attribute: theory -> src -> attribute |
22 |
val attribute_i: theory -> src -> attribute |
|
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:
16934
diff
changeset
|
26 |
(('c * 'a list) * ('d * 'a list) list) list -> |
18905 | 27 |
(('c * 'att list) * ('d * 'att list) list) list |
18734 | 28 |
val crude_closure: Context.proof -> src -> src |
29 |
val add_attributes: (bstring * (src -> attribute) * string) list -> theory -> theory |
|
18636 | 30 |
val thm: Context.generic * Args.T list -> thm * (Context.generic * Args.T list) |
31 |
val thms: Context.generic * Args.T list -> thm list * (Context.generic * Args.T list) |
|
18998 | 32 |
val multi_thm: Context.generic * Args.T list -> thm list * (Context.generic * Args.T list) |
18734 | 33 |
val syntax: (Context.generic * Args.T list -> |
34 |
attribute * (Context.generic * Args.T list)) -> src -> attribute |
|
35 |
val no_args: attribute -> src -> attribute |
|
36 |
val add_del_args: attribute -> attribute -> src -> attribute |
|
37 |
val internal: attribute -> src |
|
5823 | 38 |
end; |
39 |
||
40 |
structure Attrib: ATTRIB = |
|
41 |
struct |
|
42 |
||
15703 | 43 |
type src = Args.src; |
44 |
||
5823 | 45 |
|
18636 | 46 |
|
18734 | 47 |
(** named attributes **) |
18636 | 48 |
|
18734 | 49 |
(* theory data *) |
5823 | 50 |
|
16458 | 51 |
structure AttributesData = TheoryDataFun |
52 |
(struct |
|
5823 | 53 |
val name = "Isar/attributes"; |
18734 | 54 |
type T = (((src -> attribute) * string) * stamp) NameSpace.table; |
5823 | 55 |
|
16344 | 56 |
val empty = NameSpace.empty_table; |
6546 | 57 |
val copy = I; |
16458 | 58 |
val extend = I; |
5823 | 59 |
|
17496 | 60 |
fun merge _ tables = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUPS dups => |
16344 | 61 |
error ("Attempt to merge different versions of attribute(s) " ^ commas_quote dups); |
5823 | 62 |
|
16344 | 63 |
fun print _ attribs = |
5823 | 64 |
let |
65 |
fun prt_attr (name, ((_, comment), _)) = Pretty.block |
|
6846 | 66 |
[Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; |
5823 | 67 |
in |
16344 | 68 |
[Pretty.big_list "attributes:" (map prt_attr (NameSpace.extern_table attribs))] |
9216 | 69 |
|> Pretty.chunks |> Pretty.writeln |
5823 | 70 |
end; |
16458 | 71 |
end); |
5823 | 72 |
|
18708 | 73 |
val _ = Context.add_setup AttributesData.init; |
5823 | 74 |
val print_attributes = AttributesData.print; |
7611 | 75 |
|
5823 | 76 |
|
15703 | 77 |
(* interning *) |
78 |
||
16458 | 79 |
val intern = NameSpace.intern o #1 o AttributesData.get; |
15703 | 80 |
val intern_src = Args.map_name o intern; |
81 |
||
82 |
||
18734 | 83 |
(* get attributes *) |
5823 | 84 |
|
5912 | 85 |
exception ATTRIB_FAIL of (string * Position.T) * exn; |
86 |
||
18734 | 87 |
fun attribute_i thy = |
5823 | 88 |
let |
16344 | 89 |
val attrs = #2 (AttributesData.get thy); |
5879 | 90 |
fun attr src = |
16344 | 91 |
let val ((name, _), pos) = Args.dest_src src in |
17412 | 92 |
(case Symtab.lookup attrs name of |
15531 | 93 |
NONE => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos) |
18734 | 94 |
| SOME ((att, _), _) => transform_failure (curry ATTRIB_FAIL (name, pos)) (att src)) |
5823 | 95 |
end; |
96 |
in attr end; |
|
97 |
||
18734 | 98 |
fun attribute thy = attribute_i thy o intern_src thy; |
18636 | 99 |
|
5823 | 100 |
|
17105
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
101 |
(* attributed declarations *) |
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
102 |
|
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
103 |
fun map_specs f = map (apfst (apsnd (map f))); |
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
104 |
fun map_facts f = map (apfst (apsnd (map f)) o apsnd (map (apsnd (map f)))); |
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
105 |
|
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
106 |
|
15703 | 107 |
(* crude_closure *) |
108 |
||
109 |
(*Produce closure without knowing facts in advance! The following |
|
18734 | 110 |
works reasonably well for attribute parsers that do not peek at the |
111 |
thm structure.*) |
|
15703 | 112 |
|
113 |
fun crude_closure ctxt src = |
|
18734 | 114 |
(try (fn () => attribute_i (ProofContext.theory_of ctxt) src |
115 |
(Context.Proof ctxt, Drule.asm_rl)) (); |
|
15703 | 116 |
Args.closure src); |
117 |
||
118 |
||
5823 | 119 |
(* add_attributes *) |
120 |
||
121 |
fun add_attributes raw_attrs thy = |
|
122 |
let |
|
18734 | 123 |
val new_attrs = |
124 |
raw_attrs |> map (fn (name, att, comment) => (name, ((att, comment), stamp ()))); |
|
16458 | 125 |
fun add attrs = NameSpace.extend_table (Sign.naming_of thy) (attrs, new_attrs) |
16344 | 126 |
handle Symtab.DUPS dups => |
127 |
error ("Duplicate declaration of attributes(s) " ^ commas_quote dups); |
|
128 |
in AttributesData.map add thy end; |
|
5823 | 129 |
|
5879 | 130 |
(*implicit version*) |
131 |
fun Attribute name att cmt = Context.>> (add_attributes [(name, att, cmt)]); |
|
5823 | 132 |
|
5879 | 133 |
|
134 |
||
135 |
(** attribute parsers **) |
|
136 |
||
137 |
(* tags *) |
|
5823 | 138 |
|
5879 | 139 |
fun tag x = Scan.lift (Args.name -- Scan.repeat Args.name) x; |
140 |
||
141 |
||
142 |
(* theorems *) |
|
143 |
||
18636 | 144 |
local |
145 |
||
18998 | 146 |
val get_thms = Context.cases PureThy.get_thms ProofContext.get_thms; |
147 |
||
148 |
fun gen_thm pick = Scan.depend (fn st => |
|
149 |
(Scan.ahead Args.alt_name -- Args.named_fact (get_thms st o Fact) |
|
18037 | 150 |
>> (fn (s, fact) => ("", Fact s, fact)) || |
18998 | 151 |
Scan.ahead Args.name -- Args.named_fact (get_thms st o Name) -- Args.thm_sel |
18037 | 152 |
>> (fn ((name, fact), sel) => (name, NameSelection (name, sel), fact)) || |
18998 | 153 |
Scan.ahead Args.name -- Args.named_fact (get_thms st o Name) |
18037 | 154 |
>> (fn (name, fact) => (name, Name name, fact))) -- |
18998 | 155 |
Args.opt_attribs (intern (Context.theory_of st)) |
18037 | 156 |
>> (fn ((name, thmref, fact), srcs) => |
15703 | 157 |
let |
16498 | 158 |
val ths = PureThy.select_thm thmref fact; |
18998 | 159 |
val atts = map (attribute_i (Context.theory_of st)) srcs; |
160 |
val (st', ths') = foldl_map (Library.apply atts) (st, ths); |
|
15703 | 161 |
in (st', pick name ths') end)); |
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15117
diff
changeset
|
162 |
|
18636 | 163 |
in |
164 |
||
18998 | 165 |
val thm = gen_thm PureThy.single_thm; |
166 |
val multi_thm = gen_thm (K I); |
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19473
diff
changeset
|
167 |
val thms = Scan.repeat multi_thm >> flat; |
18636 | 168 |
|
169 |
end; |
|
170 |
||
5823 | 171 |
|
5879 | 172 |
|
173 |
(** attribute syntax **) |
|
5823 | 174 |
|
5879 | 175 |
fun syntax scan src (st, th) = |
8282 | 176 |
let val (st', f) = Args.syntax "attribute" scan src st |
5879 | 177 |
in f (st', th) end; |
178 |
||
179 |
fun no_args x = syntax (Scan.succeed x); |
|
5823 | 180 |
|
10034 | 181 |
fun add_del_args add del x = syntax |
182 |
(Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add)) x; |
|
8633 | 183 |
|
5823 | 184 |
|
185 |
||
18636 | 186 |
(** basic attributes **) |
5823 | 187 |
|
188 |
(* tags *) |
|
189 |
||
18799 | 190 |
fun tagged x = syntax (tag >> PureThy.tag) x; |
191 |
fun untagged x = syntax (Scan.lift Args.name >> PureThy.untag) x; |
|
6772 | 192 |
|
193 |
||
18636 | 194 |
(* rule composition *) |
5879 | 195 |
|
18636 | 196 |
val COMP_att = |
197 |
syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm |
|
18734 | 198 |
>> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => Drule.compose_single (A, i, B)))); |
5879 | 199 |
|
18636 | 200 |
val THEN_att = |
201 |
syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm |
|
18734 | 202 |
>> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => A RSN (i, B)))); |
5879 | 203 |
|
18636 | 204 |
val OF_att = |
18998 | 205 |
syntax (thms >> (fn Bs => Thm.rule_attribute (fn _ => fn A => Bs MRS A))); |
5879 | 206 |
|
207 |
||
15703 | 208 |
(* read_instantiate: named instantiation of type and term variables *) |
14718 | 209 |
|
15703 | 210 |
local |
5879 | 211 |
|
15703 | 212 |
fun is_tvar (x, _) = (case Symbol.explode x of "'" :: _ => true | _ => false); |
213 |
||
214 |
fun error_var msg xi = error (msg ^ Syntax.string_of_vname xi); |
|
215 |
||
15973 | 216 |
fun the_sort sorts xi = the (sorts xi) |
15703 | 217 |
handle Option.Option => error_var "No such type variable in theorem: " xi; |
10807 | 218 |
|
15973 | 219 |
fun the_type types xi = the (types xi) |
15703 | 220 |
handle Option.Option => error_var "No such variable in theorem: " xi; |
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
221 |
|
19473 | 222 |
fun unify_types thy types (xi, u) (unifier, maxidx) = |
15703 | 223 |
let |
224 |
val T = the_type types xi; |
|
225 |
val U = Term.fastype_of u; |
|
226 |
val maxidx' = Int.max (maxidx, Int.max (#2 xi, Term.maxidx_of_term u)); |
|
227 |
in |
|
16934 | 228 |
Sign.typ_unify thy (T, U) (unifier, maxidx') |
15703 | 229 |
handle Type.TUNIFY => error_var "Incompatible type for instantiation of " xi |
230 |
end; |
|
231 |
||
232 |
fun typ_subst env = apsnd (Term.typ_subst_TVars env); |
|
233 |
fun subst env = apsnd (Term.subst_TVars env); |
|
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
234 |
|
16458 | 235 |
fun instantiate thy envT env thm = |
15703 | 236 |
let |
15798
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents:
15719
diff
changeset
|
237 |
val (_, sorts) = Drule.types_sorts thm; |
16458 | 238 |
fun prepT (a, T) = (Thm.ctyp_of thy (TVar (a, the_sort sorts a)), Thm.ctyp_of thy T); |
239 |
fun prep (xi, t) = pairself (Thm.cterm_of thy) (Var (xi, Term.fastype_of t), t); |
|
15703 | 240 |
in |
19046
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
wenzelm
parents:
18998
diff
changeset
|
241 |
Drule.instantiate (map prepT (distinct (op =) envT), |
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
wenzelm
parents:
18998
diff
changeset
|
242 |
map prep (distinct (fn ((xi, t), (yj, u)) => xi = yj andalso t aconv u) env)) thm |
15703 | 243 |
end; |
244 |
||
245 |
in |
|
246 |
||
18636 | 247 |
fun read_instantiate mixed_insts (generic, thm) = |
15703 | 248 |
let |
18636 | 249 |
val thy = Context.theory_of generic; |
250 |
val ctxt = Context.proof_of generic; |
|
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
251 |
|
17756 | 252 |
val (type_insts, term_insts) = List.partition (is_tvar o fst) (map snd mixed_insts); |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19473
diff
changeset
|
253 |
val internal_insts = term_insts |> map_filter |
15703 | 254 |
(fn (xi, Args.Term t) => SOME (xi, t) |
255 |
| (_, Args.Name _) => NONE |
|
256 |
| (xi, _) => error_var "Term argument expected for " xi); |
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19473
diff
changeset
|
257 |
val external_insts = term_insts |> map_filter |
15703 | 258 |
(fn (xi, Args.Name s) => SOME (xi, s) | _ => NONE); |
5879 | 259 |
|
15703 | 260 |
|
261 |
(* type instantiations *) |
|
5879 | 262 |
|
15703 | 263 |
val sorts = #2 (Drule.types_sorts thm); |
14718 | 264 |
|
15703 | 265 |
fun readT (xi, arg) = |
266 |
let |
|
267 |
val S = the_sort sorts xi; |
|
268 |
val T = |
|
269 |
(case arg of |
|
270 |
Args.Name s => ProofContext.read_typ ctxt s |
|
271 |
| Args.Typ T => T |
|
272 |
| _ => error_var "Type argument expected for " xi); |
|
10807 | 273 |
in |
16458 | 274 |
if Sign.of_sort thy (T, S) then (xi, T) |
15703 | 275 |
else error_var "Incompatible sort for typ instantiation of " xi |
10807 | 276 |
end; |
5879 | 277 |
|
15703 | 278 |
val type_insts' = map readT type_insts; |
16458 | 279 |
val thm' = instantiate thy type_insts' [] thm; |
15703 | 280 |
|
281 |
||
282 |
(* internal term instantiations *) |
|
283 |
||
284 |
val types' = #1 (Drule.types_sorts thm'); |
|
15798
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents:
15719
diff
changeset
|
285 |
val unifier = map (apsnd snd) (Vartab.dest (#1 |
19473 | 286 |
(fold (unify_types thy types') internal_insts (Vartab.empty, 0)))); |
15703 | 287 |
|
288 |
val type_insts'' = map (typ_subst unifier) type_insts'; |
|
289 |
val internal_insts'' = map (subst unifier) internal_insts; |
|
16458 | 290 |
val thm'' = instantiate thy unifier internal_insts'' thm'; |
15703 | 291 |
|
292 |
||
293 |
(* external term instantiations *) |
|
294 |
||
295 |
val types'' = #1 (Drule.types_sorts thm''); |
|
296 |
||
297 |
val (xs, ss) = split_list external_insts; |
|
298 |
val Ts = map (the_type types'') xs; |
|
299 |
val (ts, inferred) = ProofContext.read_termTs ctxt (K false) |
|
300 |
(K NONE) (K NONE) (Drule.add_used thm'' []) (ss ~~ Ts); |
|
301 |
||
302 |
val type_insts''' = map (typ_subst inferred) type_insts''; |
|
303 |
val internal_insts''' = map (subst inferred) internal_insts''; |
|
5879 | 304 |
|
15703 | 305 |
val external_insts''' = xs ~~ ts; |
306 |
val term_insts''' = internal_insts''' @ external_insts'''; |
|
16458 | 307 |
val thm''' = instantiate thy inferred external_insts''' thm''; |
17184 | 308 |
|
15703 | 309 |
|
310 |
(* assign internalized values *) |
|
311 |
||
17184 | 312 |
val _ = |
313 |
mixed_insts |> List.app (fn (arg, (xi, _)) => |
|
314 |
if is_tvar xi then |
|
315 |
Args.assign (SOME (Args.Typ (the (AList.lookup (op =) type_insts''' xi)))) arg |
|
316 |
else |
|
317 |
Args.assign (SOME (Args.Term (the (AList.lookup (op =) term_insts''' xi)))) arg); |
|
15703 | 318 |
|
18636 | 319 |
in (generic, thm''' |> RuleCases.save thm) end; |
15703 | 320 |
|
321 |
end; |
|
322 |
||
323 |
||
324 |
(* where: named instantiation *) |
|
325 |
||
326 |
local |
|
327 |
||
328 |
val value = |
|
329 |
Args.internal_typ >> Args.Typ || |
|
330 |
Args.internal_term >> Args.Term || |
|
331 |
Args.name >> Args.Name; |
|
332 |
||
333 |
val inst = Args.var -- (Args.$$$ "=" |-- Args.ahead -- value) |
|
334 |
>> (fn (xi, (a, v)) => (a, (xi, v))); |
|
335 |
||
336 |
in |
|
5823 | 337 |
|
18636 | 338 |
val where_att = syntax (Args.and_list (Scan.lift inst) >> read_instantiate); |
5879 | 339 |
|
15703 | 340 |
end; |
5879 | 341 |
|
15703 | 342 |
|
343 |
(* of: positional instantiation (term arguments only) *) |
|
344 |
||
345 |
local |
|
5912 | 346 |
|
18636 | 347 |
fun read_instantiate' (args, concl_args) (generic, thm) = |
15703 | 348 |
let |
349 |
fun zip_vars _ [] = [] |
|
350 |
| zip_vars (_ :: xs) ((_, NONE) :: rest) = zip_vars xs rest |
|
351 |
| zip_vars ((x, _) :: xs) ((a, SOME t) :: rest) = (a, (x, t)) :: zip_vars xs rest |
|
352 |
| zip_vars [] _ = error "More instantiations than variables in theorem"; |
|
353 |
val insts = |
|
354 |
zip_vars (Drule.vars_of_terms [Thm.prop_of thm]) args @ |
|
355 |
zip_vars (Drule.vars_of_terms [Thm.concl_of thm]) concl_args; |
|
18636 | 356 |
in read_instantiate insts (generic, thm) end; |
5912 | 357 |
|
15703 | 358 |
val value = |
359 |
Args.internal_term >> Args.Term || |
|
360 |
Args.name >> Args.Name; |
|
361 |
||
362 |
val inst = Args.ahead -- Args.maybe value; |
|
10807 | 363 |
val concl = Args.$$$ "concl" -- Args.colon; |
5912 | 364 |
|
15703 | 365 |
val insts = |
366 |
Scan.repeat (Scan.unless concl inst) -- |
|
367 |
Scan.optional (concl |-- Scan.repeat inst) []; |
|
368 |
||
369 |
in |
|
5912 | 370 |
|
18636 | 371 |
val of_att = syntax (Scan.lift insts >> read_instantiate'); |
5912 | 372 |
|
15703 | 373 |
end; |
374 |
||
5912 | 375 |
|
13782
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
376 |
(* rename_abs *) |
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
377 |
|
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
378 |
fun rename_abs src = syntax |
15703 | 379 |
(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:
13414
diff
changeset
|
380 |
|
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
381 |
|
7598 | 382 |
(* unfold / fold definitions *) |
383 |
||
18839 | 384 |
fun unfolded_syntax rule = |
18998 | 385 |
syntax (thms >> |
18872 | 386 |
(fn ths => Thm.rule_attribute (fn context => rule (Context.proof_of context) ths))); |
18839 | 387 |
|
388 |
val unfolded = unfolded_syntax LocalDefs.unfold; |
|
389 |
val folded = unfolded_syntax LocalDefs.fold; |
|
7598 | 390 |
|
391 |
||
8368 | 392 |
(* rule cases *) |
393 |
||
10528 | 394 |
fun consumes x = syntax (Scan.lift (Scan.optional Args.nat 1) >> RuleCases.consumes) x; |
8368 | 395 |
fun case_names x = syntax (Scan.lift (Scan.repeat1 Args.name) >> RuleCases.case_names) x; |
18236 | 396 |
fun case_conclusion x = |
397 |
syntax (Scan.lift (Args.name -- Scan.repeat Args.name) >> RuleCases.case_conclusion) x; |
|
8368 | 398 |
fun params x = syntax (Args.and_list1 (Scan.lift (Scan.repeat Args.name)) >> RuleCases.params) x; |
399 |
||
400 |
||
11770 | 401 |
(* rule_format *) |
402 |
||
15703 | 403 |
fun rule_format_att x = syntax (Args.mode "no_asm" |
404 |
>> (fn true => ObjectLogic.rule_format_no_asm | false => ObjectLogic.rule_format)) x; |
|
11770 | 405 |
|
406 |
||
5879 | 407 |
(* misc rules *) |
408 |
||
18734 | 409 |
fun standard x = no_args (Thm.rule_attribute (K Drule.standard)) x; |
410 |
fun elim_format x = no_args (Thm.rule_attribute (K Tactic.make_elim)) x; |
|
411 |
fun no_vars x = no_args (Thm.rule_attribute (K (#1 o Drule.freeze_thaw))) x; |
|
412 |
fun eta_long x = no_args (Thm.rule_attribute (K (Drule.fconv_rule Drule.eta_long_conversion))) x; |
|
5879 | 413 |
|
414 |
||
15703 | 415 |
(* internal attribute *) |
416 |
||
18734 | 417 |
fun internal att = Args.src (("Pure.attribute", [Args.mk_attribute att]), Position.none); |
15703 | 418 |
|
18636 | 419 |
fun internal_att x = syntax (Scan.lift Args.internal_attribute) x; |
15703 | 420 |
|
421 |
||
18636 | 422 |
(* theory setup *) |
5823 | 423 |
|
18636 | 424 |
val _ = Context.add_setup |
18708 | 425 |
(add_attributes |
18734 | 426 |
[("tagged", tagged, "tagged theorem"), |
427 |
("untagged", untagged, "untagged theorem"), |
|
428 |
("COMP", COMP_att, "direct composition with rules (no lifting)"), |
|
429 |
("THEN", THEN_att, "resolution with rule"), |
|
430 |
("OF", OF_att, "rule applied to facts"), |
|
431 |
("where", where_att, "named instantiation of theorem"), |
|
432 |
("of", of_att, "rule applied to terms"), |
|
433 |
("rename_abs", rename_abs, "rename bound variables in abstractions"), |
|
434 |
("unfolded", unfolded, "unfolded definitions"), |
|
435 |
("folded", folded, "folded definitions"), |
|
436 |
("standard", standard, "result put into standard form"), |
|
437 |
("elim_format", elim_format, "destruct rule turned into elimination rule format"), |
|
438 |
("no_vars", no_vars, "frozen schematic vars"), |
|
439 |
("eta_long", eta_long, "put theorem into eta long beta normal form"), |
|
440 |
("consumes", consumes, "number of consumed facts"), |
|
441 |
("case_names", case_names, "named rule cases"), |
|
442 |
("case_conclusion", case_conclusion, "named conclusion of rule cases"), |
|
443 |
("params", params, "named rule parameters"), |
|
444 |
("atomize", no_args ObjectLogic.declare_atomize, "declaration of atomize rule"), |
|
445 |
("rulify", no_args ObjectLogic.declare_rulify, "declaration of rulify rule"), |
|
446 |
("rule_format", rule_format_att, "result put into standard rule format"), |
|
18839 | 447 |
("defn", add_del_args LocalDefs.defn_add LocalDefs.defn_del, |
448 |
"declaration of definitional transformations"), |
|
18734 | 449 |
("attribute", internal_att, "internal attribute")]); |
5823 | 450 |
|
451 |
end; |
|
452 |
||
453 |
structure BasicAttrib: BASIC_ATTRIB = Attrib; |
|
454 |
open BasicAttrib; |