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
|
5879
|
11 |
val Attribute: bstring -> (Args.src -> theory attribute) * (Args.src -> Proof.context attribute)
|
|
12 |
-> string -> unit
|
5823
|
13 |
end;
|
|
14 |
|
|
15 |
signature ATTRIB =
|
|
16 |
sig
|
|
17 |
include BASIC_ATTRIB
|
5912
|
18 |
exception ATTRIB_FAIL of (string * Position.T) * exn
|
5823
|
19 |
val global_attribute: theory -> Args.src -> theory attribute
|
|
20 |
val local_attribute: theory -> Args.src -> Proof.context attribute
|
5912
|
21 |
val local_attribute': Proof.context -> Args.src -> Proof.context attribute
|
5823
|
22 |
val add_attributes: (bstring * ((Args.src -> theory attribute) *
|
|
23 |
(Args.src -> Proof.context attribute)) * string) list -> theory -> theory
|
6091
|
24 |
val global_thm: theory * Args.T list -> thm * (theory * Args.T list)
|
|
25 |
val global_thms: theory * Args.T list -> thm list * (theory * Args.T list)
|
|
26 |
val global_thmss: theory * Args.T list -> thm list * (theory * Args.T list)
|
|
27 |
val local_thm: Proof.context * Args.T list -> thm * (Proof.context * Args.T list)
|
|
28 |
val local_thms: Proof.context * Args.T list -> thm list * (Proof.context * Args.T list)
|
|
29 |
val local_thmss: Proof.context * Args.T list -> thm list * (Proof.context * Args.T list)
|
5879
|
30 |
val syntax: ('a * Args.T list -> 'a attribute * ('a * Args.T list)) -> Args.src -> 'a attribute
|
5823
|
31 |
val no_args: 'a attribute -> Args.src -> 'a attribute
|
|
32 |
val setup: (theory -> theory) list
|
|
33 |
end;
|
|
34 |
|
|
35 |
structure Attrib: ATTRIB =
|
|
36 |
struct
|
|
37 |
|
|
38 |
|
|
39 |
(** attributes theory data **)
|
|
40 |
|
|
41 |
(* data kind 'Isar/attributes' *)
|
|
42 |
|
|
43 |
structure AttributesDataArgs =
|
|
44 |
struct
|
|
45 |
val name = "Isar/attributes";
|
|
46 |
type T =
|
|
47 |
{space: NameSpace.T,
|
|
48 |
attrs:
|
|
49 |
((((Args.src -> theory attribute) * (Args.src -> Proof.context attribute))
|
|
50 |
* string) * stamp) Symtab.table};
|
|
51 |
|
|
52 |
val empty = {space = NameSpace.empty, attrs = Symtab.empty};
|
6546
|
53 |
val copy = I;
|
5823
|
54 |
val prep_ext = I;
|
|
55 |
|
|
56 |
fun merge ({space = space1, attrs = attrs1}, {space = space2, attrs = attrs2}) =
|
|
57 |
{space = NameSpace.merge (space1, space2),
|
|
58 |
attrs = Symtab.merge eq_snd (attrs1, attrs2) handle Symtab.DUPS dups =>
|
|
59 |
error ("Attempt to merge different versions of attributes " ^ commas_quote dups)};
|
|
60 |
|
|
61 |
fun print _ ({space, attrs}) =
|
|
62 |
let
|
|
63 |
fun prt_attr (name, ((_, comment), _)) = Pretty.block
|
6846
|
64 |
[Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
|
5823
|
65 |
in
|
|
66 |
Pretty.writeln (Display.pretty_name_space ("attribute name space", space));
|
6846
|
67 |
Pretty.writeln (Pretty.big_list "attributes:"
|
|
68 |
(map prt_attr (NameSpace.cond_extern_table space attrs)))
|
5823
|
69 |
end;
|
|
70 |
end;
|
|
71 |
|
|
72 |
structure AttributesData = TheoryDataFun(AttributesDataArgs);
|
|
73 |
val print_attributes = AttributesData.print;
|
|
74 |
|
|
75 |
|
|
76 |
(* get global / local attributes *)
|
|
77 |
|
5912
|
78 |
exception ATTRIB_FAIL of (string * Position.T) * exn;
|
|
79 |
|
5823
|
80 |
fun gen_attribute which thy =
|
|
81 |
let
|
|
82 |
val {space, attrs} = AttributesData.get thy;
|
|
83 |
|
5879
|
84 |
fun attr src =
|
|
85 |
let
|
|
86 |
val ((raw_name, _), pos) = Args.dest_src src;
|
|
87 |
val name = NameSpace.intern space raw_name;
|
|
88 |
in
|
5823
|
89 |
(case Symtab.lookup (attrs, name) of
|
|
90 |
None => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
|
5912
|
91 |
| Some ((p, _), _) => transform_failure (curry ATTRIB_FAIL (name, pos)) (which p src))
|
5823
|
92 |
end;
|
|
93 |
in attr end;
|
|
94 |
|
|
95 |
val global_attribute = gen_attribute fst;
|
|
96 |
val local_attribute = gen_attribute snd;
|
5879
|
97 |
val local_attribute' = local_attribute o ProofContext.theory_of;
|
5823
|
98 |
|
|
99 |
|
|
100 |
(* add_attributes *)
|
|
101 |
|
|
102 |
fun add_attributes raw_attrs thy =
|
|
103 |
let
|
|
104 |
val full = Sign.full_name (Theory.sign_of thy);
|
|
105 |
val new_attrs =
|
|
106 |
map (fn (name, (f, g), comment) => (full name, (((f, g), comment), stamp ()))) raw_attrs;
|
|
107 |
|
|
108 |
val {space, attrs} = AttributesData.get thy;
|
|
109 |
val space' = NameSpace.extend (space, map fst new_attrs);
|
|
110 |
val attrs' = Symtab.extend (attrs, new_attrs) handle Symtab.DUPS dups =>
|
|
111 |
error ("Duplicate declaration of attributes(s) " ^ commas_quote dups);
|
|
112 |
in thy |> AttributesData.put {space = space', attrs = attrs'} end;
|
|
113 |
|
5879
|
114 |
(*implicit version*)
|
|
115 |
fun Attribute name att cmt = Context.>> (add_attributes [(name, att, cmt)]);
|
5823
|
116 |
|
5879
|
117 |
|
|
118 |
|
|
119 |
(** attribute parsers **)
|
|
120 |
|
|
121 |
(* tags *)
|
5823
|
122 |
|
5879
|
123 |
fun tag x = Scan.lift (Args.name -- Scan.repeat Args.name) x;
|
|
124 |
|
|
125 |
|
|
126 |
(* theorems *)
|
|
127 |
|
|
128 |
fun gen_thm get attrib app =
|
|
129 |
Scan.depend (fn st => Args.name -- Args.opt_attribs >>
|
|
130 |
(fn (name, srcs) => app ((st, get st name), map (attrib st) srcs)));
|
5823
|
131 |
|
6091
|
132 |
val global_thm = gen_thm PureThy.get_thm global_attribute Thm.apply_attributes;
|
|
133 |
val global_thms = gen_thm PureThy.get_thms global_attribute Thm.applys_attributes;
|
5879
|
134 |
val global_thmss = Scan.repeat global_thms >> flat;
|
|
135 |
|
6091
|
136 |
val local_thm = gen_thm ProofContext.get_thm local_attribute' Thm.apply_attributes;
|
|
137 |
val local_thms = gen_thm ProofContext.get_thms local_attribute' Thm.applys_attributes;
|
5879
|
138 |
val local_thmss = Scan.repeat local_thms >> flat;
|
|
139 |
|
5823
|
140 |
|
5879
|
141 |
|
|
142 |
(** attribute syntax **)
|
5823
|
143 |
|
5879
|
144 |
fun syntax scan src (st, th) =
|
|
145 |
let val (st', f) = Args.syntax "attribute" scan st src
|
|
146 |
in f (st', th) end;
|
|
147 |
|
|
148 |
fun no_args x = syntax (Scan.succeed x);
|
5823
|
149 |
|
|
150 |
|
|
151 |
|
|
152 |
(** Pure attributes **)
|
|
153 |
|
|
154 |
(* tags *)
|
|
155 |
|
6091
|
156 |
fun gen_tag x = syntax (tag >> Drule.tag) x;
|
|
157 |
fun gen_untag x = syntax (tag >> Drule.untag) x;
|
5823
|
158 |
|
|
159 |
|
5879
|
160 |
(* transfer *)
|
|
161 |
|
6091
|
162 |
fun gen_transfer theory_of = no_args (Drule.rule_attribute (fn st => Thm.transfer (theory_of st)));
|
5879
|
163 |
|
|
164 |
val global_transfer = gen_transfer I;
|
|
165 |
val local_transfer = gen_transfer ProofContext.theory_of;
|
|
166 |
|
|
167 |
|
6772
|
168 |
(* COMP *)
|
|
169 |
|
|
170 |
fun comp B (x, A) = (x, A COMP B);
|
|
171 |
|
|
172 |
fun gen_COMP thm = syntax (thm >> comp);
|
|
173 |
val global_COMP = gen_COMP global_thm;
|
|
174 |
val local_COMP = gen_COMP local_thm;
|
|
175 |
|
|
176 |
|
5879
|
177 |
(* RS *)
|
|
178 |
|
6091
|
179 |
fun resolve (i, B) (x, A) = (x, A RSN (i, B));
|
5879
|
180 |
|
|
181 |
fun gen_RS thm = syntax (Scan.lift (Scan.optional Args.nat 1) -- thm >> resolve);
|
|
182 |
val global_RS = gen_RS global_thm;
|
|
183 |
val local_RS = gen_RS local_thm;
|
|
184 |
|
|
185 |
|
|
186 |
(* APP *)
|
|
187 |
|
6091
|
188 |
fun apply Bs (x, A) = (x, Bs MRS A);
|
5879
|
189 |
|
|
190 |
val global_APP = syntax (global_thmss >> apply);
|
|
191 |
val local_APP = syntax (local_thmss >> apply);
|
|
192 |
|
|
193 |
|
5912
|
194 |
(* where: named instantiations *)
|
5879
|
195 |
|
5912
|
196 |
fun read_instantiate context_of insts x thm =
|
5879
|
197 |
let
|
|
198 |
val ctxt = context_of x;
|
|
199 |
val sign = ProofContext.sign_of ctxt;
|
|
200 |
|
5912
|
201 |
val vars = Drule.vars_of thm;
|
5879
|
202 |
fun get_typ xi =
|
|
203 |
(case assoc (vars, xi) of
|
|
204 |
Some T => T
|
|
205 |
| None => error ("No such variable in theorem: " ^ Syntax.string_of_vname xi));
|
|
206 |
|
5912
|
207 |
val (xs, ss) = Library.split_list insts;
|
5879
|
208 |
val Ts = map get_typ xs;
|
|
209 |
|
|
210 |
val (ts, envT) = ProofContext.read_termTs (ctxt |> ProofContext.declare_thm thm) (ss ~~ Ts);
|
|
211 |
val cenvT = map (apsnd (Thm.ctyp_of sign)) envT;
|
5912
|
212 |
val cenv =
|
|
213 |
map (fn (xi, t) => pairself (Thm.cterm_of sign) (Var (xi, fastype_of t), t))
|
|
214 |
(gen_distinct (fn ((x1, t1), (x2, t2)) => x1 = x2 andalso t1 aconv t2) (xs ~~ ts));
|
5879
|
215 |
in Thm.instantiate (cenvT, cenv) thm end;
|
|
216 |
|
6448
|
217 |
fun insts x = Args.and_list (Scan.lift (Args.var --| Args.$$$ "=" -- Args.name)) x;
|
5879
|
218 |
|
6091
|
219 |
fun gen_where context_of = syntax (insts >> (Drule.rule_attribute o read_instantiate context_of));
|
5823
|
220 |
|
5879
|
221 |
val global_where = gen_where ProofContext.init;
|
|
222 |
val local_where = gen_where I;
|
|
223 |
|
|
224 |
|
5912
|
225 |
(* with: positional instantiations *)
|
|
226 |
|
|
227 |
fun read_instantiate' context_of (args, concl_args) x thm =
|
|
228 |
let
|
|
229 |
fun zip_vars _ [] = []
|
|
230 |
| zip_vars (_ :: xs) (None :: opt_ts) = zip_vars xs opt_ts
|
|
231 |
| zip_vars ((x, _) :: xs) (Some t :: opt_ts) = (x, t) :: zip_vars xs opt_ts
|
|
232 |
| zip_vars [] _ = error "More instantiations than variables in theorem";
|
|
233 |
val insts =
|
|
234 |
zip_vars (Drule.vars_of_terms [#prop (Thm.rep_thm thm)]) args @
|
|
235 |
zip_vars (Drule.vars_of_terms [Thm.concl_of thm]) concl_args;
|
|
236 |
in read_instantiate context_of insts x thm end;
|
|
237 |
|
|
238 |
val concl = Args.$$$ "concl" -- Args.$$$ ":";
|
|
239 |
val inst_arg = Scan.unless concl (Args.$$$ "_" >> K None || Args.name >> Some);
|
|
240 |
val inst_args = Scan.repeat inst_arg;
|
|
241 |
fun insts' x = Scan.lift (inst_args -- Scan.optional (concl |-- Args.!!! inst_args) []) x;
|
|
242 |
|
6091
|
243 |
fun gen_with context_of = syntax (insts' >> (Drule.rule_attribute o read_instantiate' context_of));
|
5912
|
244 |
|
|
245 |
val global_with = gen_with ProofContext.init;
|
|
246 |
val local_with = gen_with I;
|
|
247 |
|
|
248 |
|
5879
|
249 |
(* misc rules *)
|
|
250 |
|
6091
|
251 |
fun standard x = no_args (Drule.rule_attribute (K Drule.standard)) x;
|
|
252 |
fun elimify x = no_args (Drule.rule_attribute (K Tactic.make_elim)) x;
|
5879
|
253 |
|
|
254 |
|
|
255 |
|
|
256 |
(** theory setup **)
|
5823
|
257 |
|
|
258 |
(* pure_attributes *)
|
|
259 |
|
|
260 |
val pure_attributes =
|
|
261 |
[("tag", (gen_tag, gen_tag), "tag theorem"),
|
|
262 |
("untag", (gen_untag, gen_untag), "untag theorem"),
|
6772
|
263 |
("COMP", (global_COMP, local_COMP), "compose rules (no lifting)"),
|
5879
|
264 |
("RS", (global_RS, local_RS), "resolve with rule"),
|
|
265 |
("APP", (global_APP, local_APP), "resolve rule with"),
|
5912
|
266 |
("where", (global_where, local_where), "named instantiation of theorem"),
|
|
267 |
("with", (global_with, local_with), "positional instantiation of theorem"),
|
5879
|
268 |
("standard", (standard, standard), "put theorem into standard form"),
|
|
269 |
("elimify", (elimify, elimify), "turn destruct rule into elimination rule"),
|
|
270 |
("transfer", (global_transfer, local_transfer), "transfer theorem to this theory")];
|
5823
|
271 |
|
|
272 |
|
5879
|
273 |
(* setup *)
|
5823
|
274 |
|
|
275 |
val setup = [AttributesData.init, add_attributes pure_attributes];
|
|
276 |
|
|
277 |
|
|
278 |
end;
|
|
279 |
|
|
280 |
|
|
281 |
structure BasicAttrib: BASIC_ATTRIB = Attrib;
|
|
282 |
open BasicAttrib;
|