src/Pure/Isar/attrib.ML
author wenzelm
Wed, 18 Mar 2009 22:41:14 +0100
changeset 30578 368e26dfba69
parent 30543 2ca69af4422a
child 30766 3bc78fbb9f57
permissions -rw-r--r--
more precise type Symbol_Pos.text;
     1 (*  Title:      Pure/Isar/attrib.ML
     2     Author:     Markus Wenzel, TU Muenchen
     3 
     4 Symbolic representation of attributes -- with name and syntax.
     5 *)
     6 
     7 signature ATTRIB =
     8 sig
     9   type src = Args.src
    10   type binding = binding * src list
    11   val empty_binding: binding
    12   val print_attributes: theory -> unit
    13   val intern: theory -> xstring -> string
    14   val intern_src: theory -> src -> src
    15   val pretty_attribs: Proof.context -> src list -> Pretty.T list
    16   val defined: theory -> string -> bool
    17   val attribute: theory -> src -> attribute
    18   val attribute_i: theory -> src -> attribute
    19   val eval_thms: Proof.context -> (Facts.ref * src list) list -> thm list
    20   val map_specs: ('a -> 'att) ->
    21     (('c * 'a list) * 'd) list -> (('c * 'att list) * 'd) list
    22   val map_facts: ('a -> 'att) ->
    23     (('c * 'a list) * ('d * 'a list) list) list ->
    24     (('c * 'att list) * ('d * 'att list) list) list
    25   val crude_closure: Proof.context -> src -> src
    26   val add_attributes: (bstring * (src -> attribute) * string) list -> theory -> theory
    27   val syntax: attribute context_parser -> src -> attribute
    28   val setup: Binding.binding -> attribute context_parser -> string -> theory -> theory
    29   val attribute_setup: bstring * Position.T -> Symbol_Pos.text * Position.T -> string ->
    30     theory -> theory
    31   val no_args: attribute -> src -> attribute
    32   val add_del: attribute -> attribute -> attribute context_parser
    33   val add_del_args: attribute -> attribute -> src -> attribute
    34   val thm_sel: Facts.interval list parser
    35   val thm: thm context_parser
    36   val thms: thm list context_parser
    37   val multi_thm: thm list context_parser
    38   val print_configs: Proof.context -> unit
    39   val internal: (morphism -> attribute) -> src
    40   val register_config: Config.value Config.T -> theory -> theory
    41   val config_bool: bstring -> bool -> bool Config.T * (theory -> theory)
    42   val config_int: bstring -> int -> int Config.T * (theory -> theory)
    43   val config_string: bstring -> string -> string Config.T * (theory -> theory)
    44   val config_bool_global: bstring -> bool -> bool Config.T * (theory -> theory)
    45   val config_int_global: bstring -> int -> int Config.T * (theory -> theory)
    46   val config_string_global: bstring -> string -> string Config.T * (theory -> theory)
    47 end;
    48 
    49 structure Attrib: ATTRIB =
    50 struct
    51 
    52 structure T = OuterLex;
    53 structure P = OuterParse;
    54 
    55 
    56 (* source and bindings *)
    57 
    58 type src = Args.src;
    59 
    60 type binding = binding * src list;
    61 val empty_binding: binding = (Binding.empty, []);
    62 
    63 
    64 
    65 (** named attributes **)
    66 
    67 (* theory data *)
    68 
    69 structure Attributes = TheoryDataFun
    70 (
    71   type T = (((src -> attribute) * string) * stamp) NameSpace.table;
    72   val empty = NameSpace.empty_table;
    73   val copy = I;
    74   val extend = I;
    75   fun merge _ tables : T = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUP dup =>
    76     error ("Attempt to merge different versions of attribute " ^ quote dup);
    77 );
    78 
    79 fun print_attributes thy =
    80   let
    81     val attribs = Attributes.get thy;
    82     fun prt_attr (name, ((_, comment), _)) = Pretty.block
    83       [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
    84   in
    85     [Pretty.big_list "attributes:" (map prt_attr (NameSpace.extern_table attribs))]
    86     |> Pretty.chunks |> Pretty.writeln
    87   end;
    88 
    89 
    90 (* name space *)
    91 
    92 val intern = NameSpace.intern o #1 o Attributes.get;
    93 val intern_src = Args.map_name o intern;
    94 
    95 val extern = NameSpace.extern o #1 o Attributes.get o ProofContext.theory_of;
    96 
    97 
    98 (* pretty printing *)
    99 
   100 fun pretty_attribs _ [] = []
   101   | pretty_attribs ctxt srcs =
   102       [Pretty.enclose "[" "]"
   103         (Pretty.commas (map (Args.pretty_src ctxt o Args.map_name (extern ctxt)) srcs))];
   104 
   105 
   106 (* get attributes *)
   107 
   108 val defined = Symtab.defined o #2 o Attributes.get;
   109 
   110 fun attribute_i thy =
   111   let
   112     val attrs = #2 (Attributes.get thy);
   113     fun attr src =
   114       let val ((name, _), pos) = Args.dest_src src in
   115         (case Symtab.lookup attrs name of
   116           NONE => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
   117         | SOME ((att, _), _) => (Position.report (Markup.attribute name) pos; att src))
   118       end;
   119   in attr end;
   120 
   121 fun attribute thy = attribute_i thy o intern_src thy;
   122 
   123 fun eval_thms ctxt args = ProofContext.note_thmss Thm.theoremK
   124     [(Thm.empty_binding, map (apsnd (map (attribute (ProofContext.theory_of ctxt)))) args)] ctxt
   125   |> fst |> maps snd;
   126 
   127 
   128 (* attributed declarations *)
   129 
   130 fun map_specs f = map (apfst (apsnd (map f)));
   131 fun map_facts f = map (apfst (apsnd (map f)) o apsnd (map (apsnd (map f))));
   132 
   133 
   134 (* crude_closure *)
   135 
   136 (*Produce closure without knowing facts in advance! The following
   137   works reasonably well for attribute parsers that do not peek at the
   138   thm structure.*)
   139 
   140 fun crude_closure ctxt src =
   141  (try (fn () => attribute_i (ProofContext.theory_of ctxt) src
   142     (Context.Proof ctxt, Drule.asm_rl)) ();
   143   Args.closure src);
   144 
   145 
   146 (* add_attributes *)
   147 
   148 fun add_attributes raw_attrs thy =
   149   let
   150     val new_attrs =
   151       raw_attrs |> map (fn (name, att, comment) => (Binding.name name, ((att, comment), stamp ())));
   152     fun add attrs = fold (snd oo NameSpace.define (Sign.naming_of thy)) new_attrs attrs
   153       handle Symtab.DUP dup => error ("Duplicate declaration of attributes " ^ quote dup);
   154   in Attributes.map add thy end;
   155 
   156 
   157 (* attribute setup *)
   158 
   159 fun syntax scan src (context, th) =
   160   let val (f: attribute, context') = Args.syntax "attribute" scan src context
   161   in f (context', th) end;
   162 
   163 fun setup name scan comment = add_attributes [(Binding.name_of name, syntax scan, comment)];
   164 
   165 fun attribute_setup name (txt, pos) cmt =
   166   Context.theory_map (ML_Context.expression pos
   167     "val (name, scan, comment): binding * attribute context_parser * string"
   168     "Context.map_theory (Attrib.setup name scan comment)"
   169     ("(" ^ ML_Syntax.make_binding name ^ ", " ^ txt ^ ", " ^ ML_Syntax.print_string cmt ^ ")"));
   170 
   171 
   172 (* basic syntax *)
   173 
   174 fun no_args x = syntax (Scan.succeed x);
   175 
   176 fun add_del add del = (Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add));
   177 fun add_del_args add del = syntax (add_del add del);
   178 
   179 
   180 
   181 (** parsing attributed theorems **)
   182 
   183 val thm_sel = P.$$$ "(" |-- P.list1
   184  (P.nat --| P.minus -- P.nat >> Facts.FromTo ||
   185   P.nat --| P.minus >> Facts.From ||
   186   P.nat >> Facts.Single) --| P.$$$ ")";
   187 
   188 local
   189 
   190 val fact_name = Args.internal_fact >> K "<fact>" || Args.name;
   191 
   192 fun gen_thm pick = Scan.depend (fn context =>
   193   let
   194     val thy = Context.theory_of context;
   195     val get = Context.cases (PureThy.get_fact context) ProofContext.get_fact context;
   196     val get_fact = get o Facts.Fact;
   197     fun get_named pos name = get (Facts.Named ((name, pos), NONE));
   198   in
   199     P.$$$ "[" |-- Args.attribs (intern thy) --| P.$$$ "]" >> (fn srcs =>
   200       let
   201         val atts = map (attribute_i thy) srcs;
   202         val (context', th') = Library.apply atts (context, Drule.dummy_thm);
   203       in (context', pick "" [th']) end)
   204     ||
   205     (Scan.ahead Args.alt_name -- Args.named_fact get_fact
   206       >> (fn (s, fact) => ("", Facts.Fact s, fact)) ||
   207      Scan.ahead (P.position fact_name) :|-- (fn (name, pos) =>
   208       Args.named_fact (get_named pos) -- Scan.option thm_sel
   209         >> (fn (fact, sel) => (name, Facts.Named ((name, pos), sel), fact))))
   210     -- Args.opt_attribs (intern thy) >> (fn ((name, thmref, fact), srcs) =>
   211       let
   212         val ths = Facts.select thmref fact;
   213         val atts = map (attribute_i thy) srcs;
   214         val (context', ths') = Library.foldl_map (Library.apply atts) (context, ths);
   215       in (context', pick name ths') end)
   216   end);
   217 
   218 in
   219 
   220 val thm = gen_thm Facts.the_single;
   221 val multi_thm = gen_thm (K I);
   222 val thms = Scan.repeat multi_thm >> flat;
   223 
   224 end;
   225 
   226 
   227 
   228 (** basic attributes **)
   229 
   230 (* internal *)
   231 
   232 fun internal att = Args.src (("Pure.attribute", [T.mk_attribute att]), Position.none);
   233 
   234 val internal_att =
   235   syntax (Scan.lift Args.internal_attribute >> Morphism.form);
   236 
   237 
   238 (* tags *)
   239 
   240 val tagged = syntax (Scan.lift (Args.name -- Args.name) >> Thm.tag);
   241 val untagged = syntax (Scan.lift Args.name >> Thm.untag);
   242 
   243 val kind = syntax (Scan.lift Args.name >> Thm.kind);
   244 
   245 
   246 (* rule composition *)
   247 
   248 val COMP_att =
   249   syntax (Scan.lift (Scan.optional (Args.bracks P.nat) 1) -- thm
   250     >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => Drule.compose_single (A, i, B))));
   251 
   252 val THEN_att =
   253   syntax (Scan.lift (Scan.optional (Args.bracks P.nat) 1) -- thm
   254     >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => A RSN (i, B))));
   255 
   256 val OF_att =
   257   syntax (thms >> (fn Bs => Thm.rule_attribute (fn _ => fn A => Bs MRS A)));
   258 
   259 
   260 (* rename_abs *)
   261 
   262 val rename_abs = syntax
   263   (Scan.lift (Scan.repeat (Args.maybe Args.name) >> (apsnd o Drule.rename_bvars')));
   264 
   265 
   266 (* unfold / fold definitions *)
   267 
   268 fun unfolded_syntax rule =
   269   syntax (thms >>
   270     (fn ths => Thm.rule_attribute (fn context => rule (Context.proof_of context) ths)));
   271 
   272 val unfolded = unfolded_syntax LocalDefs.unfold;
   273 val folded = unfolded_syntax LocalDefs.fold;
   274 
   275 
   276 (* rule cases *)
   277 
   278 val consumes = syntax (Scan.lift (Scan.optional P.nat 1) >> RuleCases.consumes);
   279 val case_names = syntax (Scan.lift (Scan.repeat1 Args.name) >> RuleCases.case_names);
   280 val case_conclusion =
   281   syntax (Scan.lift (Args.name -- Scan.repeat Args.name) >> RuleCases.case_conclusion);
   282 val params = syntax (Scan.lift (P.and_list1 (Scan.repeat Args.name)) >> RuleCases.params);
   283 
   284 
   285 (* rule format *)
   286 
   287 val rule_format = syntax (Args.mode "no_asm"
   288   >> (fn true => ObjectLogic.rule_format_no_asm | false => ObjectLogic.rule_format));
   289 
   290 val elim_format = no_args (Thm.rule_attribute (K Tactic.make_elim));
   291 
   292 
   293 (* misc rules *)
   294 
   295 val standard = no_args (Thm.rule_attribute (K Drule.standard));
   296 
   297 val no_vars = no_args (Thm.rule_attribute (fn context => fn th =>
   298   let
   299     val ctxt = Variable.set_body false (Context.proof_of context);
   300     val ((_, [th']), _) = Variable.import_thms true [th] ctxt;
   301   in th' end));
   302 
   303 val eta_long =
   304   no_args (Thm.rule_attribute (K (Conv.fconv_rule Drule.eta_long_conversion)));
   305 
   306 val rotated = syntax
   307   (Scan.lift (Scan.optional P.int 1) >> (fn n => Thm.rule_attribute (K (rotate_prems n))));
   308 
   309 val abs_def = no_args (Thm.rule_attribute (K Drule.abs_def));
   310 
   311 
   312 (* theory setup *)
   313 
   314 val _ = Context.>> (Context.map_theory
   315  (add_attributes
   316    [("attribute", internal_att, "internal attribute"),
   317     ("tagged", tagged, "tagged theorem"),
   318     ("untagged", untagged, "untagged theorem"),
   319     ("kind", kind, "theorem kind"),
   320     ("COMP", COMP_att, "direct composition with rules (no lifting)"),
   321     ("THEN", THEN_att, "resolution with rule"),
   322     ("OF", OF_att, "rule applied to facts"),
   323     ("rename_abs", rename_abs, "rename bound variables in abstractions"),
   324     ("unfolded", unfolded, "unfolded definitions"),
   325     ("folded", folded, "folded definitions"),
   326     ("standard", standard, "result put into standard form"),
   327     ("elim_format", elim_format, "destruct rule turned into elimination rule format"),
   328     ("no_vars", no_vars, "frozen schematic vars"),
   329     ("eta_long", eta_long, "put theorem into eta long beta normal form"),
   330     ("consumes", consumes, "number of consumed facts"),
   331     ("case_names", case_names, "named rule cases"),
   332     ("case_conclusion", case_conclusion, "named conclusion of rule cases"),
   333     ("params", params, "named rule parameters"),
   334     ("atomize", no_args ObjectLogic.declare_atomize, "declaration of atomize rule"),
   335     ("rulify", no_args ObjectLogic.declare_rulify, "declaration of rulify rule"),
   336     ("rule_format", rule_format, "result put into standard rule format"),
   337     ("rotated", rotated, "rotated theorem premises"),
   338     ("defn", add_del_args LocalDefs.defn_add LocalDefs.defn_del,
   339       "declaration of definitional transformations"),
   340     ("abs_def", abs_def, "abstract over free variables of a definition")]));
   341 
   342 
   343 
   344 (** configuration options **)
   345 
   346 (* naming *)
   347 
   348 structure Configs = TheoryDataFun
   349 (
   350   type T = Config.value Config.T Symtab.table;
   351   val empty = Symtab.empty;
   352   val copy = I;
   353   val extend = I;
   354   fun merge _ = Symtab.merge (K true);
   355 );
   356 
   357 fun print_configs ctxt =
   358   let
   359     val thy = ProofContext.theory_of ctxt;
   360     fun prt (name, config) =
   361       let val value = Config.get ctxt config in
   362         Pretty.block [Pretty.str (name ^ ": " ^ Config.print_type value ^ " ="), Pretty.brk 1,
   363           Pretty.str (Config.print_value value)]
   364       end;
   365     val configs = NameSpace.extern_table (#1 (Attributes.get thy), Configs.get thy);
   366   in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end;
   367 
   368 
   369 (* concrete syntax *)
   370 
   371 local
   372 
   373 val equals = P.$$$ "=";
   374 
   375 fun scan_value (Config.Bool _) =
   376       equals -- Args.$$$ "false" >> K (Config.Bool false) ||
   377       equals -- Args.$$$ "true" >> K (Config.Bool true) ||
   378       Scan.succeed (Config.Bool true)
   379   | scan_value (Config.Int _) = equals |-- P.int >> Config.Int
   380   | scan_value (Config.String _) = equals |-- Args.name >> Config.String;
   381 
   382 fun scan_config thy config =
   383   let val config_type = Config.get_thy thy config
   384   in scan_value config_type >> (K o Thm.declaration_attribute o K o Config.put_generic config) end;
   385 
   386 in
   387 
   388 fun register_config config thy =
   389   let
   390     val bname = Config.name_of config;
   391     val name = Sign.full_bname thy bname;
   392   in
   393     thy
   394     |> add_attributes [(bname, syntax (Scan.lift (scan_config thy config) >> Morphism.form),
   395       "configuration option")]
   396     |> Configs.map (Symtab.update (name, config))
   397   end;
   398 
   399 fun declare_config make coerce global name default =
   400   let
   401     val config_value = Config.declare global name (make default);
   402     val config = coerce config_value;
   403   in (config, register_config config_value) end;
   404 
   405 val config_bool   = declare_config Config.Bool Config.bool false;
   406 val config_int    = declare_config Config.Int Config.int false;
   407 val config_string = declare_config Config.String Config.string false;
   408 
   409 val config_bool_global   = declare_config Config.Bool Config.bool true;
   410 val config_int_global    = declare_config Config.Int Config.int true;
   411 val config_string_global = declare_config Config.String Config.string true;
   412 
   413 end;
   414 
   415 
   416 (* theory setup *)
   417 
   418 val _ = Context.>> (Context.map_theory
   419  (register_config Unify.trace_bound_value #>
   420   register_config Unify.search_bound_value #>
   421   register_config Unify.trace_simp_value #>
   422   register_config Unify.trace_types_value #>
   423   register_config MetaSimplifier.simp_depth_limit_value));
   424 
   425 end;