author | wenzelm |
Wed, 23 Nov 2005 18:52:03 +0100 | |
changeset 18236 | dd445f5cb28e |
parent 18037 | 1095d2213b9d |
child 18243 | 1287b15f27ef |
permissions | -rw-r--r-- |
5823 | 1 |
(* Title: Pure/Isar/attrib.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
5 |
Symbolic theorem attributes. |
|
6 |
*) |
|
7 |
||
8 |
signature BASIC_ATTRIB = |
|
9 |
sig |
|
10 |
val print_attributes: theory -> unit |
|
16458 | 11 |
val Attribute: bstring -> |
12 |
(Args.src -> theory attribute) * (Args.src -> ProofContext.context attribute) -> |
|
13 |
string -> unit |
|
5823 | 14 |
end; |
15 |
||
16 |
signature ATTRIB = |
|
17 |
sig |
|
18 |
include BASIC_ATTRIB |
|
15703 | 19 |
type src |
5912 | 20 |
exception ATTRIB_FAIL of (string * Position.T) * exn |
16458 | 21 |
val intern: theory -> xstring -> string |
22 |
val intern_src: theory -> src -> src |
|
15703 | 23 |
val global_attribute_i: theory -> src -> theory attribute |
24 |
val global_attribute: theory -> src -> theory attribute |
|
25 |
val local_attribute_i: theory -> src -> ProofContext.context attribute |
|
26 |
val local_attribute: theory -> src -> ProofContext.context attribute |
|
27 |
val context_attribute_i: ProofContext.context -> Args.src -> ProofContext.context attribute |
|
28 |
val context_attribute: ProofContext.context -> Args.src -> ProofContext.context attribute |
|
7673
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
29 |
val undef_global_attribute: theory attribute |
15703 | 30 |
val undef_local_attribute: ProofContext.context attribute |
17105
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
31 |
val map_specs: ('a -> 'b attribute) -> |
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
32 |
(('c * 'a list) * 'd) list -> (('c * 'b attribute list) * 'd) list |
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
33 |
val map_facts: ('a -> 'b attribute) -> |
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
34 |
(('c * 'a list) * ('d * 'a list) list) list -> |
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
35 |
(('c * 'b attribute list) * ('d * 'b attribute list) list) list |
15703 | 36 |
val crude_closure: ProofContext.context -> src -> src |
37 |
val add_attributes: (bstring * ((src -> theory attribute) * |
|
38 |
(src -> ProofContext.context attribute)) * string) list -> theory -> theory |
|
6091 | 39 |
val global_thm: theory * Args.T list -> thm * (theory * Args.T list) |
40 |
val global_thms: theory * Args.T list -> thm list * (theory * Args.T list) |
|
41 |
val global_thmss: theory * Args.T list -> thm list * (theory * Args.T list) |
|
16458 | 42 |
val local_thm: ProofContext.context * Args.T list -> |
43 |
thm * (ProofContext.context * Args.T list) |
|
44 |
val local_thms: ProofContext.context * Args.T list -> |
|
45 |
thm list * (ProofContext.context * Args.T list) |
|
46 |
val local_thmss: ProofContext.context * Args.T list -> |
|
47 |
thm list * (ProofContext.context * Args.T list) |
|
15703 | 48 |
val syntax: ('a * Args.T list -> 'a attribute * ('a * Args.T list)) -> src -> 'a attribute |
49 |
val no_args: 'a attribute -> src -> 'a attribute |
|
50 |
val add_del_args: 'a attribute -> 'a attribute -> src -> 'a attribute |
|
51 |
val attribute: (theory attribute * ProofContext.context attribute) -> src |
|
5823 | 52 |
end; |
53 |
||
54 |
structure Attrib: ATTRIB = |
|
55 |
struct |
|
56 |
||
15703 | 57 |
type src = Args.src; |
58 |
||
5823 | 59 |
|
60 |
(** attributes theory data **) |
|
61 |
||
16458 | 62 |
(* datatype attributes *) |
5823 | 63 |
|
16458 | 64 |
structure AttributesData = TheoryDataFun |
65 |
(struct |
|
5823 | 66 |
val name = "Isar/attributes"; |
67 |
type T = |
|
16344 | 68 |
((((src -> theory attribute) * (src -> ProofContext.context attribute)) |
69 |
* string) * stamp) NameSpace.table; |
|
5823 | 70 |
|
16344 | 71 |
val empty = NameSpace.empty_table; |
6546 | 72 |
val copy = I; |
16458 | 73 |
val extend = I; |
5823 | 74 |
|
17496 | 75 |
fun merge _ tables = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUPS dups => |
16344 | 76 |
error ("Attempt to merge different versions of attribute(s) " ^ commas_quote dups); |
5823 | 77 |
|
16344 | 78 |
fun print _ attribs = |
5823 | 79 |
let |
80 |
fun prt_attr (name, ((_, comment), _)) = Pretty.block |
|
6846 | 81 |
[Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; |
5823 | 82 |
in |
16344 | 83 |
[Pretty.big_list "attributes:" (map prt_attr (NameSpace.extern_table attribs))] |
9216 | 84 |
|> Pretty.chunks |> Pretty.writeln |
5823 | 85 |
end; |
16458 | 86 |
end); |
5823 | 87 |
|
15801 | 88 |
val _ = Context.add_setup [AttributesData.init]; |
5823 | 89 |
val print_attributes = AttributesData.print; |
7611 | 90 |
|
5823 | 91 |
|
15703 | 92 |
(* interning *) |
93 |
||
16458 | 94 |
val intern = NameSpace.intern o #1 o AttributesData.get; |
15703 | 95 |
val intern_src = Args.map_name o intern; |
96 |
||
97 |
||
5823 | 98 |
(* get global / local attributes *) |
99 |
||
5912 | 100 |
exception ATTRIB_FAIL of (string * Position.T) * exn; |
101 |
||
5823 | 102 |
fun gen_attribute which thy = |
103 |
let |
|
16344 | 104 |
val attrs = #2 (AttributesData.get thy); |
5879 | 105 |
fun attr src = |
16344 | 106 |
let val ((name, _), pos) = Args.dest_src src in |
17412 | 107 |
(case Symtab.lookup attrs name of |
15531 | 108 |
NONE => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos) |
109 |
| SOME ((p, _), _) => transform_failure (curry ATTRIB_FAIL (name, pos)) (which p src)) |
|
5823 | 110 |
end; |
111 |
in attr end; |
|
112 |
||
15703 | 113 |
val global_attribute_i = gen_attribute fst; |
16458 | 114 |
fun global_attribute thy = global_attribute_i thy o intern_src thy; |
15703 | 115 |
|
116 |
val local_attribute_i = gen_attribute snd; |
|
16458 | 117 |
fun local_attribute thy = local_attribute_i thy o intern_src thy; |
15703 | 118 |
|
119 |
val context_attribute_i = local_attribute_i o ProofContext.theory_of; |
|
120 |
val context_attribute = local_attribute o ProofContext.theory_of; |
|
5823 | 121 |
|
7673
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
122 |
val undef_global_attribute: theory attribute = |
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
123 |
fn _ => error "attribute undefined in theory context"; |
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
124 |
|
15703 | 125 |
val undef_local_attribute: ProofContext.context attribute = |
7673
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
126 |
fn _ => error "attribute undefined in proof context"; |
b8e7fa177d62
added undef_global_attribute, undef_local_attribute;
wenzelm
parents:
7668
diff
changeset
|
127 |
|
5823 | 128 |
|
17105
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
129 |
(* attributed declarations *) |
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
130 |
|
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
131 |
fun map_specs f = map (apfst (apsnd (map f))); |
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
132 |
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
|
133 |
|
1b07a176a880
added map_specs/facts operators (from locale.ML);
wenzelm
parents:
16934
diff
changeset
|
134 |
|
15703 | 135 |
(* crude_closure *) |
136 |
||
137 |
(*Produce closure without knowing facts in advance! The following |
|
138 |
should work reasonably well for attribute parsers that do not peek |
|
139 |
at the thm structure.*) |
|
140 |
||
141 |
fun crude_closure ctxt src = |
|
142 |
(try (transform_error (fn () => context_attribute_i ctxt src (ctxt, Drule.asm_rl))) (); |
|
143 |
Args.closure src); |
|
144 |
||
145 |
||
5823 | 146 |
(* add_attributes *) |
147 |
||
148 |
fun add_attributes raw_attrs thy = |
|
149 |
let |
|
16141 | 150 |
val new_attrs = raw_attrs |> map (fn (name, (f, g), comment) => |
16344 | 151 |
(name, (((f, g), comment), stamp ()))); |
16458 | 152 |
fun add attrs = NameSpace.extend_table (Sign.naming_of thy) (attrs, new_attrs) |
16344 | 153 |
handle Symtab.DUPS dups => |
154 |
error ("Duplicate declaration of attributes(s) " ^ commas_quote dups); |
|
155 |
in AttributesData.map add thy end; |
|
5823 | 156 |
|
5879 | 157 |
(*implicit version*) |
158 |
fun Attribute name att cmt = Context.>> (add_attributes [(name, att, cmt)]); |
|
5823 | 159 |
|
5879 | 160 |
|
161 |
||
162 |
(** attribute parsers **) |
|
163 |
||
164 |
(* tags *) |
|
5823 | 165 |
|
5879 | 166 |
fun tag x = Scan.lift (Args.name -- Scan.repeat Args.name) x; |
167 |
||
168 |
||
169 |
(* theorems *) |
|
170 |
||
15703 | 171 |
fun gen_thm theory_of attrib get pick = Scan.depend (fn st => |
18037 | 172 |
(Scan.ahead Args.alt_name -- Args.named_fact (get st o Fact) |
173 |
>> (fn (s, fact) => ("", Fact s, fact)) || |
|
174 |
Scan.ahead Args.name -- Args.named_fact (get st o Name) -- Args.thm_sel |
|
175 |
>> (fn ((name, fact), sel) => (name, NameSelection (name, sel), fact)) || |
|
176 |
Scan.ahead Args.name -- Args.named_fact (get st o Name) |
|
177 |
>> (fn (name, fact) => (name, Name name, fact))) -- |
|
178 |
Args.opt_attribs (intern (theory_of st)) |
|
179 |
>> (fn ((name, thmref, fact), srcs) => |
|
15703 | 180 |
let |
16498 | 181 |
val ths = PureThy.select_thm thmref fact; |
15703 | 182 |
val atts = map (attrib (theory_of st)) srcs; |
183 |
val (st', ths') = Thm.applys_attributes ((st, ths), atts); |
|
184 |
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
|
185 |
|
15703 | 186 |
val global_thm = gen_thm I global_attribute_i PureThy.get_thms PureThy.single_thm; |
187 |
val global_thms = gen_thm I global_attribute_i PureThy.get_thms (K I); |
|
15570 | 188 |
val global_thmss = Scan.repeat global_thms >> List.concat; |
5879 | 189 |
|
15703 | 190 |
val local_thm = |
191 |
gen_thm ProofContext.theory_of local_attribute_i ProofContext.get_thms PureThy.single_thm; |
|
192 |
val local_thms = |
|
193 |
gen_thm ProofContext.theory_of local_attribute_i ProofContext.get_thms (K I); |
|
15570 | 194 |
val local_thmss = Scan.repeat local_thms >> List.concat; |
5879 | 195 |
|
5823 | 196 |
|
5879 | 197 |
|
198 |
(** attribute syntax **) |
|
5823 | 199 |
|
5879 | 200 |
fun syntax scan src (st, th) = |
8282 | 201 |
let val (st', f) = Args.syntax "attribute" scan src st |
5879 | 202 |
in f (st', th) end; |
203 |
||
204 |
fun no_args x = syntax (Scan.succeed x); |
|
5823 | 205 |
|
10034 | 206 |
fun add_del_args add del x = syntax |
207 |
(Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add)) x; |
|
8633 | 208 |
|
5823 | 209 |
|
210 |
||
211 |
(** Pure attributes **) |
|
212 |
||
213 |
(* tags *) |
|
214 |
||
9902 | 215 |
fun gen_tagged x = syntax (tag >> Drule.tag) x; |
216 |
fun gen_untagged x = syntax (Scan.lift Args.name >> Drule.untag) x; |
|
5823 | 217 |
|
218 |
||
6772 | 219 |
(* COMP *) |
220 |
||
6948 | 221 |
fun comp (i, B) (x, A) = (x, Drule.compose_single (A, i, B)); |
6772 | 222 |
|
10151 | 223 |
fun gen_COMP thm = syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm >> comp); |
9902 | 224 |
val COMP_global = gen_COMP global_thm; |
225 |
val COMP_local = gen_COMP local_thm; |
|
6772 | 226 |
|
227 |
||
15117 | 228 |
(* THEN, which corresponds to RS *) |
5879 | 229 |
|
6091 | 230 |
fun resolve (i, B) (x, A) = (x, A RSN (i, B)); |
5879 | 231 |
|
15117 | 232 |
fun gen_THEN thm = syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm >> resolve); |
233 |
val THEN_global = gen_THEN global_thm; |
|
234 |
val THEN_local = gen_THEN local_thm; |
|
5879 | 235 |
|
236 |
||
9902 | 237 |
(* OF *) |
5879 | 238 |
|
6091 | 239 |
fun apply Bs (x, A) = (x, Bs MRS A); |
5879 | 240 |
|
9902 | 241 |
val OF_global = syntax (global_thmss >> apply); |
242 |
val OF_local = syntax (local_thmss >> apply); |
|
5879 | 243 |
|
244 |
||
15703 | 245 |
(* read_instantiate: named instantiation of type and term variables *) |
14718 | 246 |
|
15703 | 247 |
local |
5879 | 248 |
|
15703 | 249 |
fun is_tvar (x, _) = (case Symbol.explode x of "'" :: _ => true | _ => false); |
250 |
||
251 |
fun error_var msg xi = error (msg ^ Syntax.string_of_vname xi); |
|
252 |
||
15973 | 253 |
fun the_sort sorts xi = the (sorts xi) |
15703 | 254 |
handle Option.Option => error_var "No such type variable in theorem: " xi; |
10807 | 255 |
|
15973 | 256 |
fun the_type types xi = the (types xi) |
15703 | 257 |
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
|
258 |
|
16458 | 259 |
fun unify_types thy types ((unifier, maxidx), (xi, u)) = |
15703 | 260 |
let |
261 |
val T = the_type types xi; |
|
262 |
val U = Term.fastype_of u; |
|
263 |
val maxidx' = Int.max (maxidx, Int.max (#2 xi, Term.maxidx_of_term u)); |
|
264 |
in |
|
16934 | 265 |
Sign.typ_unify thy (T, U) (unifier, maxidx') |
15703 | 266 |
handle Type.TUNIFY => error_var "Incompatible type for instantiation of " xi |
267 |
end; |
|
268 |
||
269 |
fun typ_subst env = apsnd (Term.typ_subst_TVars env); |
|
270 |
fun subst env = apsnd (Term.subst_TVars env); |
|
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
271 |
|
16458 | 272 |
fun instantiate thy envT env thm = |
15703 | 273 |
let |
15798
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents:
15719
diff
changeset
|
274 |
val (_, sorts) = Drule.types_sorts thm; |
16458 | 275 |
fun prepT (a, T) = (Thm.ctyp_of thy (TVar (a, the_sort sorts a)), Thm.ctyp_of thy T); |
276 |
fun prep (xi, t) = pairself (Thm.cterm_of thy) (Var (xi, Term.fastype_of t), t); |
|
15703 | 277 |
in |
278 |
Drule.instantiate (map prepT (distinct envT), |
|
15798
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents:
15719
diff
changeset
|
279 |
map prep (gen_distinct (fn ((xi, t), (yj, u)) => xi = yj andalso t aconv u) env)) thm |
15703 | 280 |
end; |
281 |
||
282 |
in |
|
283 |
||
284 |
fun read_instantiate init mixed_insts (context, thm) = |
|
285 |
let |
|
286 |
val ctxt = init context; |
|
16458 | 287 |
val thy = ProofContext.theory_of ctxt; |
14287
f630017ed01c
Isar: where attribute supports instantiation of type variables.
ballarin
parents:
14257
diff
changeset
|
288 |
|
17756 | 289 |
val (type_insts, term_insts) = List.partition (is_tvar o fst) (map snd mixed_insts); |
15703 | 290 |
val internal_insts = term_insts |> List.mapPartial |
291 |
(fn (xi, Args.Term t) => SOME (xi, t) |
|
292 |
| (_, Args.Name _) => NONE |
|
293 |
| (xi, _) => error_var "Term argument expected for " xi); |
|
294 |
val external_insts = term_insts |> List.mapPartial |
|
295 |
(fn (xi, Args.Name s) => SOME (xi, s) | _ => NONE); |
|
5879 | 296 |
|
15703 | 297 |
|
298 |
(* type instantiations *) |
|
5879 | 299 |
|
15703 | 300 |
val sorts = #2 (Drule.types_sorts thm); |
14718 | 301 |
|
15703 | 302 |
fun readT (xi, arg) = |
303 |
let |
|
304 |
val S = the_sort sorts xi; |
|
305 |
val T = |
|
306 |
(case arg of |
|
307 |
Args.Name s => ProofContext.read_typ ctxt s |
|
308 |
| Args.Typ T => T |
|
309 |
| _ => error_var "Type argument expected for " xi); |
|
10807 | 310 |
in |
16458 | 311 |
if Sign.of_sort thy (T, S) then (xi, T) |
15703 | 312 |
else error_var "Incompatible sort for typ instantiation of " xi |
10807 | 313 |
end; |
5879 | 314 |
|
15703 | 315 |
val type_insts' = map readT type_insts; |
16458 | 316 |
val thm' = instantiate thy type_insts' [] thm; |
15703 | 317 |
|
318 |
||
319 |
(* internal term instantiations *) |
|
320 |
||
321 |
val types' = #1 (Drule.types_sorts thm'); |
|
15798
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents:
15719
diff
changeset
|
322 |
val unifier = map (apsnd snd) (Vartab.dest (#1 |
16458 | 323 |
(Library.foldl (unify_types thy types') ((Vartab.empty, 0), internal_insts)))); |
15703 | 324 |
|
325 |
val type_insts'' = map (typ_subst unifier) type_insts'; |
|
326 |
val internal_insts'' = map (subst unifier) internal_insts; |
|
16458 | 327 |
val thm'' = instantiate thy unifier internal_insts'' thm'; |
15703 | 328 |
|
329 |
||
330 |
(* external term instantiations *) |
|
331 |
||
332 |
val types'' = #1 (Drule.types_sorts thm''); |
|
333 |
||
334 |
val (xs, ss) = split_list external_insts; |
|
335 |
val Ts = map (the_type types'') xs; |
|
336 |
val (ts, inferred) = ProofContext.read_termTs ctxt (K false) |
|
337 |
(K NONE) (K NONE) (Drule.add_used thm'' []) (ss ~~ Ts); |
|
338 |
||
339 |
val type_insts''' = map (typ_subst inferred) type_insts''; |
|
340 |
val internal_insts''' = map (subst inferred) internal_insts''; |
|
5879 | 341 |
|
15703 | 342 |
val external_insts''' = xs ~~ ts; |
343 |
val term_insts''' = internal_insts''' @ external_insts'''; |
|
16458 | 344 |
val thm''' = instantiate thy inferred external_insts''' thm''; |
17184 | 345 |
|
15703 | 346 |
|
347 |
(* assign internalized values *) |
|
348 |
||
17184 | 349 |
val _ = |
350 |
mixed_insts |> List.app (fn (arg, (xi, _)) => |
|
351 |
if is_tvar xi then |
|
352 |
Args.assign (SOME (Args.Typ (the (AList.lookup (op =) type_insts''' xi)))) arg |
|
353 |
else |
|
354 |
Args.assign (SOME (Args.Term (the (AList.lookup (op =) term_insts''' xi)))) arg); |
|
15703 | 355 |
|
17184 | 356 |
in (context, thm''' |> RuleCases.save thm) end; |
15703 | 357 |
|
358 |
end; |
|
359 |
||
360 |
||
361 |
(* where: named instantiation *) |
|
362 |
||
363 |
local |
|
364 |
||
365 |
val value = |
|
366 |
Args.internal_typ >> Args.Typ || |
|
367 |
Args.internal_term >> Args.Term || |
|
368 |
Args.name >> Args.Name; |
|
369 |
||
370 |
val inst = Args.var -- (Args.$$$ "=" |-- Args.ahead -- value) |
|
371 |
>> (fn (xi, (a, v)) => (a, (xi, v))); |
|
372 |
||
373 |
fun gen_where init = |
|
374 |
syntax (Args.and_list (Scan.lift inst) >> read_instantiate init); |
|
375 |
||
376 |
in |
|
5823 | 377 |
|
9902 | 378 |
val where_global = gen_where ProofContext.init; |
379 |
val where_local = gen_where I; |
|
5879 | 380 |
|
15703 | 381 |
end; |
5879 | 382 |
|
15703 | 383 |
|
384 |
(* of: positional instantiation (term arguments only) *) |
|
385 |
||
386 |
local |
|
5912 | 387 |
|
15703 | 388 |
fun read_instantiate' init (args, concl_args) (context, thm) = |
389 |
let |
|
390 |
fun zip_vars _ [] = [] |
|
391 |
| zip_vars (_ :: xs) ((_, NONE) :: rest) = zip_vars xs rest |
|
392 |
| zip_vars ((x, _) :: xs) ((a, SOME t) :: rest) = (a, (x, t)) :: zip_vars xs rest |
|
393 |
| zip_vars [] _ = error "More instantiations than variables in theorem"; |
|
394 |
val insts = |
|
395 |
zip_vars (Drule.vars_of_terms [Thm.prop_of thm]) args @ |
|
396 |
zip_vars (Drule.vars_of_terms [Thm.concl_of thm]) concl_args; |
|
397 |
in |
|
398 |
read_instantiate init insts (context, thm) |
|
399 |
end; |
|
5912 | 400 |
|
15703 | 401 |
val value = |
402 |
Args.internal_term >> Args.Term || |
|
403 |
Args.name >> Args.Name; |
|
404 |
||
405 |
val inst = Args.ahead -- Args.maybe value; |
|
10807 | 406 |
val concl = Args.$$$ "concl" -- Args.colon; |
5912 | 407 |
|
15703 | 408 |
val insts = |
409 |
Scan.repeat (Scan.unless concl inst) -- |
|
410 |
Scan.optional (concl |-- Scan.repeat inst) []; |
|
411 |
||
412 |
fun gen_of init = syntax (Scan.lift insts >> read_instantiate' init); |
|
413 |
||
414 |
in |
|
5912 | 415 |
|
9902 | 416 |
val of_global = gen_of ProofContext.init; |
417 |
val of_local = gen_of I; |
|
5912 | 418 |
|
15703 | 419 |
end; |
420 |
||
5912 | 421 |
|
13782
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
422 |
(* rename_abs *) |
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
423 |
|
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
424 |
fun rename_abs src = syntax |
15703 | 425 |
(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
|
426 |
|
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
427 |
|
7598 | 428 |
(* unfold / fold definitions *) |
429 |
||
430 |
fun gen_rewrite rew defs (x, thm) = (x, rew defs thm); |
|
431 |
||
9902 | 432 |
val unfolded_global = syntax (global_thmss >> gen_rewrite Tactic.rewrite_rule); |
433 |
val unfolded_local = syntax (local_thmss >> gen_rewrite Tactic.rewrite_rule); |
|
434 |
val folded_global = syntax (global_thmss >> gen_rewrite Tactic.fold_rule); |
|
435 |
val folded_local = syntax (local_thmss >> gen_rewrite Tactic.fold_rule); |
|
7598 | 436 |
|
437 |
||
8368 | 438 |
(* rule cases *) |
439 |
||
10528 | 440 |
fun consumes x = syntax (Scan.lift (Scan.optional Args.nat 1) >> RuleCases.consumes) x; |
8368 | 441 |
fun case_names x = syntax (Scan.lift (Scan.repeat1 Args.name) >> RuleCases.case_names) x; |
18236 | 442 |
fun case_conclusion x = |
443 |
syntax (Scan.lift (Args.name -- Scan.repeat Args.name) >> RuleCases.case_conclusion) x; |
|
8368 | 444 |
fun params x = syntax (Args.and_list1 (Scan.lift (Scan.repeat Args.name)) >> RuleCases.params) x; |
445 |
||
446 |
||
11770 | 447 |
(* rule_format *) |
448 |
||
15703 | 449 |
fun rule_format_att x = syntax (Args.mode "no_asm" |
450 |
>> (fn true => ObjectLogic.rule_format_no_asm | false => ObjectLogic.rule_format)) x; |
|
11770 | 451 |
|
452 |
||
5879 | 453 |
(* misc rules *) |
454 |
||
6091 | 455 |
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
|
456 |
fun elim_format x = no_args (Drule.rule_attribute (K Tactic.make_elim)) x; |
9216 | 457 |
fun no_vars x = no_args (Drule.rule_attribute (K (#1 o Drule.freeze_thaw))) x; |
15926 | 458 |
fun eta_long x = no_args (Drule.rule_attribute (K (Drule.fconv_rule Drule.eta_long_conversion))) x; |
5879 | 459 |
|
460 |
||
13370 | 461 |
(* rule declarations *) |
462 |
||
463 |
local |
|
464 |
||
465 |
fun add_args a b c x = syntax |
|
466 |
(Scan.lift ((Args.bang >> K a || Args.query >> K c || Scan.succeed b) -- (Scan.option Args.nat)) |
|
467 |
>> (fn (f, n) => f n)) x; |
|
468 |
||
469 |
fun del_args att = syntax (Scan.lift Args.del >> K att); |
|
470 |
||
471 |
open ContextRules; |
|
472 |
||
473 |
in |
|
474 |
||
475 |
val rule_atts = |
|
476 |
[("intro", |
|
477 |
(add_args intro_bang_global intro_global intro_query_global, |
|
478 |
add_args intro_bang_local intro_local intro_query_local), |
|
479 |
"declaration of introduction rule"), |
|
480 |
("elim", |
|
481 |
(add_args elim_bang_global elim_global elim_query_global, |
|
482 |
add_args elim_bang_local elim_local elim_query_local), |
|
483 |
"declaration of elimination rule"), |
|
484 |
("dest", |
|
485 |
(add_args dest_bang_global dest_global dest_query_global, |
|
486 |
add_args dest_bang_local dest_local dest_query_local), |
|
487 |
"declaration of destruction rule"), |
|
488 |
("rule", (del_args rule_del_global, del_args rule_del_local), |
|
489 |
"remove declaration of intro/elim/dest rule")]; |
|
490 |
||
491 |
end; |
|
492 |
||
493 |
||
15703 | 494 |
(* internal attribute *) |
495 |
||
15828 | 496 |
fun attribute att = Args.src (("Pure.attribute", [Args.mk_attribute att]), Position.none); |
15703 | 497 |
|
498 |
fun attribute_global x = (syntax (Scan.lift Args.internal_attribute >> #1)) x; |
|
499 |
fun attribute_local x = (syntax (Scan.lift Args.internal_attribute >> #2)) x; |
|
500 |
||
15801 | 501 |
val _ = Context.add_setup |
15828 | 502 |
[add_attributes [("attribute", (attribute_global, attribute_local), "internal attribute")]]; |
15703 | 503 |
|
504 |
||
5823 | 505 |
(* pure_attributes *) |
506 |
||
507 |
val pure_attributes = |
|
9902 | 508 |
[("tagged", (gen_tagged, gen_tagged), "tagged theorem"), |
509 |
("untagged", (gen_untagged, gen_untagged), "untagged theorem"), |
|
510 |
("COMP", (COMP_global, COMP_local), "direct composition with rules (no lifting)"), |
|
15117 | 511 |
("THEN", (THEN_global, THEN_local), "resolution with rule"), |
9902 | 512 |
("OF", (OF_global, OF_local), "rule applied to facts"), |
513 |
("where", (where_global, where_local), "named instantiation of theorem"), |
|
514 |
("of", (of_global, of_local), "rule applied to terms"), |
|
13782
44de406a7273
Added rename_abs attribute for renaming bound variables.
berghofe
parents:
13414
diff
changeset
|
515 |
("rename_abs", (rename_abs, rename_abs), "rename bound variables in abstractions"), |
9902 | 516 |
("unfolded", (unfolded_global, unfolded_local), "unfolded definitions"), |
517 |
("folded", (folded_global, folded_local), "folded definitions"), |
|
518 |
("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
|
519 |
("elim_format", (elim_format, elim_format), "destruct rule turned into elimination rule format"), |
9902 | 520 |
("no_vars", (no_vars, no_vars), "frozen schematic vars"), |
15926 | 521 |
("eta_long", (eta_long, eta_long), "put theorem into eta long beta normal form"), |
10528 | 522 |
("consumes", (consumes, consumes), "number of consumed facts"), |
9902 | 523 |
("case_names", (case_names, case_names), "named rule cases"), |
18236 | 524 |
("case_clusion", (case_conclusion, case_conclusion), "named conclusion of rule cases"), |
9902 | 525 |
("params", (params, params), "named rule parameters"), |
11770 | 526 |
("atomize", (no_args ObjectLogic.declare_atomize, no_args undef_local_attribute), |
527 |
"declaration of atomize rule"), |
|
528 |
("rulify", (no_args ObjectLogic.declare_rulify, no_args undef_local_attribute), |
|
529 |
"declaration of rulify rule"), |
|
13370 | 530 |
("rule_format", (rule_format_att, rule_format_att), "result put into standard rule format")] @ |
531 |
rule_atts; |
|
5823 | 532 |
|
15801 | 533 |
val _ = Context.add_setup [add_attributes pure_attributes]; |
5823 | 534 |
|
535 |
||
536 |
end; |
|
537 |
||
538 |
structure BasicAttrib: BASIC_ATTRIB = Attrib; |
|
539 |
open BasicAttrib; |