author | obua |
Sun, 09 May 2004 23:04:36 +0200 | |
changeset 14722 | 8e739a6eaf11 |
parent 14718 | f52f2cf2d137 |
child 14981 | e73f8140af78 |
permissions | -rw-r--r-- |
5823 | 1 |
(* Title: Pure/Isar/attrib.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
8807 | 4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
5823 | 5 |
|
6 |
Symbolic theorem attributes. |
|
7 |
*) |
|
8 |
||
9 |
signature BASIC_ATTRIB = |
|
10 |
sig |
|
11 |
val print_attributes: theory -> unit |
|
5879 | 12 |
val Attribute: bstring -> (Args.src -> theory attribute) * (Args.src -> Proof.context attribute) |
13 |
-> string -> unit |
|
5823 | 14 |
end; |
15 |
||
16 |
signature ATTRIB = |
|
17 |
sig |
|
18 |
include BASIC_ATTRIB |
|
5912 | 19 |
exception ATTRIB_FAIL of (string * Position.T) * exn |
5823 | 20 |
val global_attribute: theory -> Args.src -> theory attribute |
21 |
val local_attribute: theory -> Args.src -> Proof.context attribute |
|
5912 | 22 |
val local_attribute': Proof.context -> Args.src -> Proof.context attribute |
7673
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
23 |
val undef_global_attribute: theory attribute |
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
24 |
val undef_local_attribute: Proof.context attribute |
5823 | 25 |
val add_attributes: (bstring * ((Args.src -> theory attribute) * |
26 |
(Args.src -> Proof.context attribute)) * string) list -> theory -> theory |
|
6091 | 27 |
val global_thm: theory * Args.T list -> thm * (theory * Args.T list) |
28 |
val global_thms: theory * Args.T list -> thm list * (theory * Args.T list) |
|
29 |
val global_thmss: theory * Args.T list -> thm list * (theory * Args.T list) |
|
30 |
val local_thm: Proof.context * Args.T list -> thm * (Proof.context * Args.T list) |
|
31 |
val local_thms: Proof.context * Args.T list -> thm list * (Proof.context * Args.T list) |
|
32 |
val local_thmss: Proof.context * Args.T list -> thm list * (Proof.context * Args.T list) |
|
5879 | 33 |
val syntax: ('a * Args.T list -> 'a attribute * ('a * Args.T list)) -> Args.src -> 'a attribute |
5823 | 34 |
val no_args: 'a attribute -> Args.src -> 'a attribute |
8633 | 35 |
val add_del_args: 'a attribute -> 'a attribute -> Args.src -> 'a attribute |
5823 | 36 |
val setup: (theory -> theory) list |
37 |
end; |
|
38 |
||
39 |
structure Attrib: ATTRIB = |
|
40 |
struct |
|
41 |
||
42 |
||
43 |
(** attributes theory data **) |
|
44 |
||
45 |
(* data kind 'Isar/attributes' *) |
|
46 |
||
47 |
structure AttributesDataArgs = |
|
48 |
struct |
|
49 |
val name = "Isar/attributes"; |
|
50 |
type T = |
|
51 |
{space: NameSpace.T, |
|
52 |
attrs: |
|
53 |
((((Args.src -> theory attribute) * (Args.src -> Proof.context attribute)) |
|
54 |
* string) * stamp) Symtab.table}; |
|
55 |
||
56 |
val empty = {space = NameSpace.empty, attrs = Symtab.empty}; |
|
6546 | 57 |
val copy = I; |
5823 | 58 |
val prep_ext = I; |
59 |
||
60 |
fun merge ({space = space1, attrs = attrs1}, {space = space2, attrs = attrs2}) = |
|
61 |
{space = NameSpace.merge (space1, space2), |
|
62 |
attrs = Symtab.merge eq_snd (attrs1, attrs2) handle Symtab.DUPS dups => |
|
63 |
error ("Attempt to merge different versions of attributes " ^ commas_quote dups)}; |
|
64 |
||
9216 | 65 |
fun print _ {space, attrs} = |
5823 | 66 |
let |
67 |
fun prt_attr (name, ((_, comment), _)) = Pretty.block |
|
6846 | 68 |
[Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; |
5823 | 69 |
in |
8720 | 70 |
[Pretty.big_list "attributes:" (map prt_attr (NameSpace.cond_extern_table space attrs))] |
9216 | 71 |
|> Pretty.chunks |> Pretty.writeln |
5823 | 72 |
end; |
73 |
end; |
|
74 |
||
75 |
structure AttributesData = TheoryDataFun(AttributesDataArgs); |
|
76 |
val print_attributes = AttributesData.print; |
|
7611 | 77 |
|
5823 | 78 |
|
79 |
(* get global / local attributes *) |
|
80 |
||
5912 | 81 |
exception ATTRIB_FAIL of (string * Position.T) * exn; |
82 |
||
5823 | 83 |
fun gen_attribute which thy = |
84 |
let |
|
85 |
val {space, attrs} = AttributesData.get thy; |
|
86 |
||
5879 | 87 |
fun attr src = |
88 |
let |
|
89 |
val ((raw_name, _), pos) = Args.dest_src src; |
|
90 |
val name = NameSpace.intern space raw_name; |
|
91 |
in |
|
5823 | 92 |
(case Symtab.lookup (attrs, name) of |
93 |
None => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos) |
|
5912 | 94 |
| Some ((p, _), _) => transform_failure (curry ATTRIB_FAIL (name, pos)) (which p src)) |
5823 | 95 |
end; |
96 |
in attr end; |
|
97 |
||
98 |
val global_attribute = gen_attribute fst; |
|
99 |
val local_attribute = gen_attribute snd; |
|
5879 | 100 |
val local_attribute' = local_attribute o ProofContext.theory_of; |
5823 | 101 |
|
7673
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
102 |
val undef_global_attribute: theory attribute = |
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
103 |
fn _ => error "attribute undefined in theory context"; |
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
104 |
|
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
105 |
val undef_local_attribute: Proof.context attribute = |
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
106 |
fn _ => error "attribute undefined in proof context"; |
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
107 |
|
5823 | 108 |
|
109 |
(* add_attributes *) |
|
110 |
||
111 |
fun add_attributes raw_attrs thy = |
|
112 |
let |
|
113 |
val full = Sign.full_name (Theory.sign_of thy); |
|
114 |
val new_attrs = |
|
115 |
map (fn (name, (f, g), comment) => (full name, (((f, g), comment), stamp ()))) raw_attrs; |
|
116 |
||
117 |
val {space, attrs} = AttributesData.get thy; |
|
118 |
val space' = NameSpace.extend (space, map fst new_attrs); |
|
119 |
val attrs' = Symtab.extend (attrs, new_attrs) handle Symtab.DUPS dups => |
|
120 |
error ("Duplicate declaration of attributes(s) " ^ commas_quote dups); |
|
121 |
in thy |> AttributesData.put {space = space', attrs = attrs'} end; |
|
122 |
||
5879 | 123 |
(*implicit version*) |
124 |
fun Attribute name att cmt = Context.>> (add_attributes [(name, att, cmt)]); |
|
5823 | 125 |
|
5879 | 126 |
|
127 |
||
128 |
(** attribute parsers **) |
|
129 |
||
130 |
(* tags *) |
|
5823 | 131 |
|
5879 | 132 |
fun tag x = Scan.lift (Args.name -- Scan.repeat Args.name) x; |
133 |
||
134 |
||
135 |
(* theorems *) |
|
136 |
||
137 |
fun gen_thm get attrib app = |
|
138 |
Scan.depend (fn st => Args.name -- Args.opt_attribs >> |
|
139 |
(fn (name, srcs) => app ((st, get st name), map (attrib st) srcs))); |
|
5823 | 140 |
|
6091 | 141 |
val global_thm = gen_thm PureThy.get_thm global_attribute Thm.apply_attributes; |
142 |
val global_thms = gen_thm PureThy.get_thms global_attribute Thm.applys_attributes; |
|
5879 | 143 |
val global_thmss = Scan.repeat global_thms >> flat; |
144 |
||
6091 | 145 |
val local_thm = gen_thm ProofContext.get_thm local_attribute' Thm.apply_attributes; |
146 |
val local_thms = gen_thm ProofContext.get_thms local_attribute' Thm.applys_attributes; |
|
5879 | 147 |
val local_thmss = Scan.repeat local_thms >> flat; |
148 |
||
5823 | 149 |
|
5879 | 150 |
|
151 |
(** attribute syntax **) |
|
5823 | 152 |
|
5879 | 153 |
fun syntax scan src (st, th) = |
8282 | 154 |
let val (st', f) = Args.syntax "attribute" scan src st |
5879 | 155 |
in f (st', th) end; |
156 |
||
157 |
fun no_args x = syntax (Scan.succeed x); |
|
5823 | 158 |
|
10034 | 159 |
fun add_del_args add del x = syntax |
160 |
(Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add)) x; |
|
8633 | 161 |
|
5823 | 162 |
|
163 |
||
164 |
(** Pure attributes **) |
|
165 |
||
166 |
(* tags *) |
|
167 |
||
9902 | 168 |
fun gen_tagged x = syntax (tag >> Drule.tag) x; |
169 |
fun gen_untagged x = syntax (Scan.lift Args.name >> Drule.untag) x; |
|
5823 | 170 |
|
171 |
||
6772 | 172 |
(* COMP *) |
173 |
||
6948 | 174 |
fun comp (i, B) (x, A) = (x, Drule.compose_single (A, i, B)); |
6772 | 175 |
|
10151 | 176 |
fun gen_COMP thm = syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm >> comp); |
9902 | 177 |
val COMP_global = gen_COMP global_thm; |
178 |
val COMP_local = gen_COMP local_thm; |
|
6772 | 179 |
|
180 |
||
5879 | 181 |
(* RS *) |
182 |
||
6091 | 183 |
fun resolve (i, B) (x, A) = (x, A RSN (i, B)); |
5879 | 184 |
|
10151 | 185 |
fun gen_RS thm = syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm >> resolve); |
9902 | 186 |
val RS_global = gen_RS global_thm; |
187 |
val RS_local = gen_RS local_thm; |
|
5879 | 188 |
|
189 |
||
9902 | 190 |
(* OF *) |
5879 | 191 |
|
6091 | 192 |
fun apply Bs (x, A) = (x, Bs MRS A); |
5879 | 193 |
|
9902 | 194 |
val OF_global = syntax (global_thmss >> apply); |
195 |
val OF_local = syntax (local_thmss >> apply); |
|
5879 | 196 |
|
197 |
||
14718 | 198 |
(* where *) |
199 |
||
200 |
(*named instantiations; permits instantiation of type and term variables*) |
|
5879 | 201 |
|
10807 | 202 |
fun read_instantiate _ [] _ thm = thm |
203 |
| read_instantiate context_of insts x thm = |
|
204 |
let |
|
205 |
val ctxt = context_of x; |
|
206 |
val sign = ProofContext.sign_of ctxt; |
|
207 |
||
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
208 |
(* Separate type and term insts, |
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
209 |
type insts must occur strictly before term insts *) |
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
210 |
fun has_type_var ((x, _), _) = (case Symbol.explode x of |
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
211 |
"'"::cs => true | cs => false); |
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
212 |
val (Tinst, tinsts) = take_prefix has_type_var insts; |
14718 | 213 |
val _ = if exists has_type_var tinsts |
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
214 |
then error |
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
215 |
"Type instantiations must occur before term instantiations." |
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
216 |
else (); |
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
217 |
|
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
218 |
val Tinsts = filter has_type_var insts; |
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
219 |
val tinsts = filter_out has_type_var insts; |
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
220 |
|
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
221 |
(* Type instantiations first *) |
14718 | 222 |
(* Process type insts: Tinsts_env *) |
223 |
fun absent xi = error |
|
224 |
("No such type variable in theorem: " ^ |
|
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
225 |
Syntax.string_of_vname xi); |
14718 | 226 |
val (rtypes, rsorts) = types_sorts thm; |
227 |
fun readT (xi, s) = |
|
228 |
let val S = case rsorts xi of Some S => S | None => absent xi; |
|
229 |
val T = ProofContext.read_typ ctxt s; |
|
230 |
in if Sign.typ_instance sign (T, TVar (xi, S)) then (xi, T) |
|
231 |
else error |
|
232 |
("Instantiation of " ^ Syntax.string_of_vname xi ^ " fails") |
|
233 |
end; |
|
234 |
val Tinsts_env = map readT Tinsts; |
|
235 |
val cenvT = map (apsnd (Thm.ctyp_of sign)) (Tinsts_env); |
|
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
236 |
val thm' = Thm.instantiate (cenvT, []) thm; |
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
237 |
(* Instantiate, but don't normalise: *) |
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
238 |
(* this happens after term insts anyway. *) |
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
239 |
|
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
240 |
(* Term instantiations *) |
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
241 |
val vars = Drule.vars_of thm'; |
10807 | 242 |
fun get_typ xi = |
243 |
(case assoc (vars, xi) of |
|
244 |
Some T => T |
|
245 |
| None => error ("No such variable in theorem: " ^ Syntax.string_of_vname xi)); |
|
5879 | 246 |
|
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
247 |
val (xs, ss) = Library.split_list tinsts; |
10807 | 248 |
val Ts = map get_typ xs; |
5879 | 249 |
|
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
250 |
val used = add_term_tvarnames (prop_of thm',[]) |
14718 | 251 |
(* Names of TVars occuring in thm'. read_def_termTs ensures |
14257
a7ef3f7588c5
Type inference bug in Isar attributes "where" and "of" fixed.
ballarin
parents:
14082
diff
changeset
|
252 |
that new TVars introduced when reading the instantiation string |
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
253 |
are distinct from those occuring in the theorem. *) |
14257
a7ef3f7588c5
Type inference bug in Isar attributes "where" and "of" fixed.
ballarin
parents:
14082
diff
changeset
|
254 |
|
14718 | 255 |
val (ts, envT) = |
256 |
ProofContext.read_termTs ctxt (K false) (K None) (K None) used (ss ~~ Ts); |
|
257 |
||
10807 | 258 |
val cenvT = map (apsnd (Thm.ctyp_of sign)) envT; |
259 |
val cenv = |
|
260 |
map (fn (xi, t) => pairself (Thm.cterm_of sign) (Var (xi, fastype_of t), t)) |
|
261 |
(gen_distinct (fn ((x1, t1), (x2, t2)) => x1 = x2 andalso t1 aconv t2) (xs ~~ ts)); |
|
262 |
in |
|
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
263 |
thm' |
10807 | 264 |
|> Drule.instantiate (cenvT, cenv) |
265 |
|> RuleCases.save thm |
|
266 |
end; |
|
5879 | 267 |
|
6448 | 268 |
fun insts x = Args.and_list (Scan.lift (Args.var --| Args.$$$ "=" -- Args.name)) x; |
5879 | 269 |
|
6091 | 270 |
fun gen_where context_of = syntax (insts >> (Drule.rule_attribute o read_instantiate context_of)); |
5823 | 271 |
|
9902 | 272 |
val where_global = gen_where ProofContext.init; |
273 |
val where_local = gen_where I; |
|
5879 | 274 |
|
275 |
||
9902 | 276 |
(* of: positional instantiations *) |
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
277 |
(* permits instantiation of term variables only *) |
5912 | 278 |
|
10807 | 279 |
fun read_instantiate' _ ([], []) _ thm = thm |
280 |
| read_instantiate' context_of (args, concl_args) x thm = |
|
281 |
let |
|
282 |
fun zip_vars _ [] = [] |
|
283 |
| zip_vars (_ :: xs) (None :: opt_ts) = zip_vars xs opt_ts |
|
284 |
| zip_vars ((x, _) :: xs) (Some t :: opt_ts) = (x, t) :: zip_vars xs opt_ts |
|
285 |
| zip_vars [] _ = error "More instantiations than variables in theorem"; |
|
286 |
val insts = |
|
12804 | 287 |
zip_vars (Drule.vars_of_terms [Thm.prop_of thm]) args @ |
10807 | 288 |
zip_vars (Drule.vars_of_terms [Thm.concl_of thm]) concl_args; |
289 |
in |
|
290 |
thm |
|
291 |
|> read_instantiate context_of insts x |
|
292 |
|> RuleCases.save thm |
|
293 |
end; |
|
5912 | 294 |
|
10807 | 295 |
val concl = Args.$$$ "concl" -- Args.colon; |
8687 | 296 |
val inst_arg = Scan.unless concl Args.name_dummy; |
5912 | 297 |
val inst_args = Scan.repeat inst_arg; |
10807 | 298 |
fun insts' x = (inst_args -- Scan.optional (concl |-- Args.!!! inst_args) []) x; |
5912 | 299 |
|
10807 | 300 |
fun gen_of context_of = |
301 |
syntax (Scan.lift insts' >> (Drule.rule_attribute o read_instantiate' context_of)); |
|
5912 | 302 |
|
9902 | 303 |
val of_global = gen_of ProofContext.init; |
304 |
val of_local = gen_of I; |
|
5912 | 305 |
|
306 |
||
13782
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
307 |
(* rename_abs *) |
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
308 |
|
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
309 |
fun rename_abs src = syntax |
14082
c69d5bf3047d
Moved function for renaming bound variables to Pure/drule.ML
berghofe
parents:
13782
diff
changeset
|
310 |
(Scan.lift (Scan.repeat Args.name_dummy >> (apsnd o Drule.rename_bvars'))) src; |
13782
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
311 |
|
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
312 |
|
7598 | 313 |
(* unfold / fold definitions *) |
314 |
||
315 |
fun gen_rewrite rew defs (x, thm) = (x, rew defs thm); |
|
316 |
||
9902 | 317 |
val unfolded_global = syntax (global_thmss >> gen_rewrite Tactic.rewrite_rule); |
318 |
val unfolded_local = syntax (local_thmss >> gen_rewrite Tactic.rewrite_rule); |
|
319 |
val folded_global = syntax (global_thmss >> gen_rewrite Tactic.fold_rule); |
|
320 |
val folded_local = syntax (local_thmss >> gen_rewrite Tactic.fold_rule); |
|
7598 | 321 |
|
322 |
||
8368 | 323 |
(* rule cases *) |
324 |
||
10528 | 325 |
fun consumes x = syntax (Scan.lift (Scan.optional Args.nat 1) >> RuleCases.consumes) x; |
8368 | 326 |
fun case_names x = syntax (Scan.lift (Scan.repeat1 Args.name) >> RuleCases.case_names) x; |
327 |
fun params x = syntax (Args.and_list1 (Scan.lift (Scan.repeat Args.name)) >> RuleCases.params) x; |
|
328 |
||
329 |
||
11770 | 330 |
(* rule_format *) |
331 |
||
332 |
fun rule_format_att x = syntax |
|
333 |
(Scan.lift (Args.parens (Args.$$$ "no_asm") |
|
334 |
>> K ObjectLogic.rule_format_no_asm || Scan.succeed ObjectLogic.rule_format)) x; |
|
335 |
||
336 |
||
5879 | 337 |
(* misc rules *) |
338 |
||
6091 | 339 |
fun standard x = no_args (Drule.rule_attribute (K Drule.standard)) x; |
9941
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents:
9902
diff
changeset
|
340 |
fun elim_format x = no_args (Drule.rule_attribute (K Tactic.make_elim)) x; |
9216 | 341 |
fun no_vars x = no_args (Drule.rule_attribute (K (#1 o Drule.freeze_thaw))) x; |
5879 | 342 |
|
343 |
||
13370 | 344 |
(* rule declarations *) |
345 |
||
346 |
local |
|
347 |
||
348 |
fun add_args a b c x = syntax |
|
349 |
(Scan.lift ((Args.bang >> K a || Args.query >> K c || Scan.succeed b) -- (Scan.option Args.nat)) |
|
350 |
>> (fn (f, n) => f n)) x; |
|
351 |
||
352 |
fun del_args att = syntax (Scan.lift Args.del >> K att); |
|
353 |
||
354 |
open ContextRules; |
|
355 |
||
356 |
in |
|
357 |
||
358 |
val rule_atts = |
|
359 |
[("intro", |
|
360 |
(add_args intro_bang_global intro_global intro_query_global, |
|
361 |
add_args intro_bang_local intro_local intro_query_local), |
|
362 |
"declaration of introduction rule"), |
|
363 |
("elim", |
|
364 |
(add_args elim_bang_global elim_global elim_query_global, |
|
365 |
add_args elim_bang_local elim_local elim_query_local), |
|
366 |
"declaration of elimination rule"), |
|
367 |
("dest", |
|
368 |
(add_args dest_bang_global dest_global dest_query_global, |
|
369 |
add_args dest_bang_local dest_local dest_query_local), |
|
370 |
"declaration of destruction rule"), |
|
371 |
("rule", (del_args rule_del_global, del_args rule_del_local), |
|
372 |
"remove declaration of intro/elim/dest rule")]; |
|
373 |
||
374 |
end; |
|
375 |
||
376 |
||
11770 | 377 |
|
5879 | 378 |
(** theory setup **) |
5823 | 379 |
|
380 |
(* pure_attributes *) |
|
381 |
||
382 |
val pure_attributes = |
|
9902 | 383 |
[("tagged", (gen_tagged, gen_tagged), "tagged theorem"), |
384 |
("untagged", (gen_untagged, gen_untagged), "untagged theorem"), |
|
385 |
("COMP", (COMP_global, COMP_local), "direct composition with rules (no lifting)"), |
|
386 |
("THEN", (RS_global, RS_local), "resolution with rule"), |
|
387 |
("OF", (OF_global, OF_local), "rule applied to facts"), |
|
388 |
("where", (where_global, where_local), "named instantiation of theorem"), |
|
389 |
("of", (of_global, of_local), "rule applied to terms"), |
|
13782
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
390 |
("rename_abs", (rename_abs, rename_abs), "rename bound variables in abstractions"), |
9902 | 391 |
("unfolded", (unfolded_global, unfolded_local), "unfolded definitions"), |
392 |
("folded", (folded_global, folded_local), "folded definitions"), |
|
393 |
("standard", (standard, standard), "result put into standard form"), |
|
9941
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents:
9902
diff
changeset
|
394 |
("elim_format", (elim_format, elim_format), "destruct rule turned into elimination rule format"), |
9902 | 395 |
("no_vars", (no_vars, no_vars), "frozen schematic vars"), |
10528 | 396 |
("consumes", (consumes, consumes), "number of consumed facts"), |
9902 | 397 |
("case_names", (case_names, case_names), "named rule cases"), |
398 |
("params", (params, params), "named rule parameters"), |
|
11770 | 399 |
("atomize", (no_args ObjectLogic.declare_atomize, no_args undef_local_attribute), |
400 |
"declaration of atomize rule"), |
|
401 |
("rulify", (no_args ObjectLogic.declare_rulify, no_args undef_local_attribute), |
|
402 |
"declaration of rulify rule"), |
|
13370 | 403 |
("rule_format", (rule_format_att, rule_format_att), "result put into standard rule format")] @ |
404 |
rule_atts; |
|
5823 | 405 |
|
406 |
||
5879 | 407 |
(* setup *) |
5823 | 408 |
|
409 |
val setup = [AttributesData.init, add_attributes pure_attributes]; |
|
410 |
||
411 |
end; |
|
412 |
||
413 |
structure BasicAttrib: BASIC_ATTRIB = Attrib; |
|
414 |
open BasicAttrib; |