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