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