src/Pure/Isar/attrib.ML
author wenzelm
Sat, 19 Nov 2011 14:31:43 +0100
changeset 46458 c94f149cdf5d
parent 46456 41a768a431a6
child 46483 a3ed5b65b85e
permissions -rw-r--r--
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
tuned;
     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 is_empty_binding: binding -> bool
    13   val print_attributes: theory -> unit
    14   val intern: theory -> xstring -> string
    15   val intern_src: theory -> src -> src
    16   val pretty_attribs: Proof.context -> src list -> Pretty.T list
    17   val defined: theory -> string -> bool
    18   val attribute: theory -> src -> attribute
    19   val attribute_i: theory -> src -> attribute
    20   val map_specs: ('a list -> 'att list) ->
    21     (('c * 'a list) * 'b) list -> (('c * 'att list) * 'b) list
    22   val map_facts: ('a list -> 'att list) ->
    23     (('c * 'a list) * ('d * 'a list) list) list ->
    24     (('c * 'att list) * ('d * 'att list) list) list
    25   val map_facts_refs: ('a list -> 'att list) -> ('b -> 'fact) ->
    26     (('c * 'a list) * ('b * 'a list) list) list ->
    27     (('c * 'att list) * ('fact * 'att list) list) list
    28   val eval_thms: Proof.context -> (Facts.ref * src list) list -> thm list
    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 partial_evaluation: Proof.context ->
    38     (binding * (thm list * Args.src list) list) list ->
    39     (binding * (thm list * Args.src list) list) list
    40   val internal: (morphism -> attribute) -> src
    41   val print_configs: Proof.context -> unit
    42   val config_bool: Binding.binding ->
    43     (Context.generic -> bool) -> bool Config.T * (theory -> theory)
    44   val config_int: Binding.binding ->
    45     (Context.generic -> int) -> int Config.T * (theory -> theory)
    46   val config_real: Binding.binding ->
    47     (Context.generic -> real) -> real Config.T * (theory -> theory)
    48   val config_string: Binding.binding ->
    49     (Context.generic -> string) -> string Config.T * (theory -> theory)
    50   val setup_config_bool: Binding.binding -> (Context.generic -> bool) -> bool Config.T
    51   val setup_config_int: Binding.binding -> (Context.generic -> int) -> int Config.T
    52   val setup_config_string: Binding.binding -> (Context.generic -> string) -> string Config.T
    53   val setup_config_real: Binding.binding -> (Context.generic -> real) -> real Config.T
    54 end;
    55 
    56 structure Attrib: ATTRIB =
    57 struct
    58 
    59 (* source and bindings *)
    60 
    61 type src = Args.src;
    62 
    63 type binding = binding * src list;
    64 
    65 val empty_binding: binding = (Binding.empty, []);
    66 fun is_empty_binding ((b, srcs): binding) = Binding.is_empty b andalso null srcs;
    67 
    68 
    69 
    70 (** named attributes **)
    71 
    72 (* theory data *)
    73 
    74 structure Attributes = Theory_Data
    75 (
    76   type T = ((src -> attribute) * string) Name_Space.table;
    77   val empty : T = Name_Space.empty_table "attribute";
    78   val extend = I;
    79   fun merge data : T = Name_Space.merge_tables data;
    80 );
    81 
    82 fun print_attributes thy =
    83   let
    84     val ctxt = Proof_Context.init_global thy;
    85     val attribs = Attributes.get thy;
    86     fun prt_attr (name, (_, "")) = Pretty.str name
    87       | prt_attr (name, (_, comment)) =
    88           Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
    89   in
    90     [Pretty.big_list "attributes:" (map prt_attr (Name_Space.extern_table ctxt attribs))]
    91     |> Pretty.chunks |> Pretty.writeln
    92   end;
    93 
    94 fun add_attribute name att comment thy = thy
    95   |> Attributes.map
    96     (Name_Space.define (Proof_Context.init_global thy) true (Sign.naming_of thy)
    97       (name, (att, comment)) #> snd);
    98 
    99 
   100 (* name space *)
   101 
   102 val intern = Name_Space.intern o #1 o Attributes.get;
   103 val intern_src = Args.map_name o intern;
   104 
   105 fun extern ctxt = Name_Space.extern ctxt (#1 (Attributes.get (Proof_Context.theory_of ctxt)));
   106 
   107 
   108 (* pretty printing *)
   109 
   110 fun pretty_attribs _ [] = []
   111   | pretty_attribs ctxt srcs =
   112       [Pretty.enum "," "[" "]" (map (Args.pretty_src ctxt o Args.map_name (extern ctxt)) srcs)];
   113 
   114 
   115 (* get attributes *)
   116 
   117 val defined = Symtab.defined o #2 o Attributes.get;
   118 
   119 fun attribute_i thy =
   120   let
   121     val (space, tab) = Attributes.get thy;
   122     fun attr src =
   123       let val ((name, _), pos) = Args.dest_src src in
   124         (case Symtab.lookup tab name of
   125           NONE => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
   126         | SOME (att, _) => (Position.report pos (Name_Space.markup space name); att src))
   127       end;
   128   in attr end;
   129 
   130 fun attribute thy = attribute_i thy o intern_src thy;
   131 
   132 
   133 (* attributed declarations *)
   134 
   135 fun map_specs f = map (apfst (apsnd f));
   136 
   137 fun map_facts f = map (apfst (apsnd f) o apsnd (map (apsnd f)));
   138 fun map_facts_refs f g = map_facts f #> map (apsnd (map (apfst g)));
   139 
   140 
   141 (* fact expressions *)
   142 
   143 fun eval_thms ctxt srcs = ctxt
   144   |> Proof_Context.note_thmss ""
   145     (map_facts_refs (map (attribute (Proof_Context.theory_of ctxt))) (Proof_Context.get_fact ctxt)
   146       [((Binding.empty, []), srcs)])
   147   |> fst |> maps snd;
   148 
   149 
   150 (* attribute setup *)
   151 
   152 fun syntax scan = Args.syntax "attribute" scan;
   153 
   154 fun setup name scan =
   155   add_attribute name
   156     (fn src => fn (ctxt, th) => let val (a, ctxt') = syntax scan src ctxt in a (ctxt', th) end);
   157 
   158 fun attribute_setup name (txt, pos) cmt =
   159   Context.theory_map (ML_Context.expression pos
   160     "val (name, scan, comment): binding * attribute context_parser * string"
   161     "Context.map_theory (Attrib.setup name scan comment)"
   162     (ML_Lex.read Position.none ("(" ^ ML_Syntax.make_binding name ^ ", ") @
   163       ML_Lex.read pos txt @
   164       ML_Lex.read Position.none (", " ^ 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 = Parse.$$$ "(" |-- Parse.list1
   176  (Parse.nat --| Parse.minus -- Parse.nat >> Facts.FromTo ||
   177   Parse.nat --| Parse.minus >> Facts.From ||
   178   Parse.nat >> Facts.Single) --| Parse.$$$ ")";
   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 (Global_Theory.get_fact context) Proof_Context.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     Parse.$$$ "[" |-- Args.attribs (intern thy) --| Parse.$$$ "]" >> (fn srcs =>
   192       let
   193         val atts = map (attribute_i thy) srcs;
   194         val (context', th') =
   195           Library.apply (map Thm.apply_attribute atts) (context, Drule.dummy_thm);
   196       in (context', pick "" [th']) end)
   197     ||
   198     (Scan.ahead Args.alt_name -- Args.named_fact get_fact
   199       >> (fn (s, fact) => ("", Facts.Fact s, fact)) ||
   200      Scan.ahead (Parse.position fact_name) :|-- (fn (name, pos) =>
   201       Args.named_fact (get_named pos) -- Scan.option thm_sel
   202         >> (fn (fact, sel) => (name, Facts.Named ((name, pos), sel), fact))))
   203     -- Args.opt_attribs (intern thy) >> (fn ((name, thmref, fact), srcs) =>
   204       let
   205         val ths = Facts.select thmref fact;
   206         val atts = map (attribute_i thy) srcs;
   207         val (context', ths') =
   208           Library.foldl_map (Library.apply (map Thm.apply_attribute atts)) (context, ths);
   209       in (context', pick name ths') end)
   210   end);
   211 
   212 in
   213 
   214 val thm = gen_thm Facts.the_single;
   215 val multi_thm = gen_thm (K I);
   216 val thms = Scan.repeat multi_thm >> flat;
   217 
   218 end;
   219 
   220 
   221 
   222 (** partial evaluation -- observing rule/declaration/mixed attributes **)
   223 
   224 local
   225 
   226 val strict_eq_thm = op = o pairself Thm.rep_thm;
   227 
   228 fun apply_att src (context, th) =
   229   let
   230     val src1 = Args.assignable src;
   231     val result = attribute_i (Context.theory_of context) src1 (context, th);
   232     val src2 = Args.closure src1;
   233   in (src2, result) end;
   234 
   235 fun err msg src =
   236   let val ((name, _), pos) = Args.dest_src src
   237   in error (msg ^ " " ^ quote name ^ Position.str_of pos) end;
   238 
   239 fun eval src ((th, dyn), (decls, context)) =
   240   (case (apply_att src (context, th), dyn) of
   241     ((_, (NONE, SOME th')), NONE) => ((th', NONE), (decls, context))
   242   | ((_, (NONE, SOME _)), SOME _) => err "Mixed dynamic attribute followed by static rule" src
   243   | ((src', (SOME context', NONE)), NONE) =>
   244       let
   245         val decls' =
   246           (case decls of
   247             [] => [(th, [src'])]
   248           | (th2, srcs2) :: rest =>
   249               if strict_eq_thm (th, th2)
   250               then ((th2, src' :: srcs2) :: rest)
   251               else (th, [src']) :: (th2, srcs2) :: rest);
   252       in ((th, NONE), (decls', context')) end
   253   | ((src', (opt_context', opt_th')), _) =>
   254       let
   255         val context' = the_default context opt_context';
   256         val th' = the_default th opt_th';
   257         val dyn' =
   258           (case dyn of
   259             NONE => SOME (th, [src'])
   260           | SOME (dyn_th, srcs) => SOME (dyn_th, src' :: srcs));
   261       in ((th', dyn'), (decls, context')) end);
   262 
   263 in
   264 
   265 fun partial_evaluation ctxt facts =
   266   (facts, Context.Proof (Context_Position.set_visible false ctxt)) |->
   267     fold_map (fn ((b, more_atts), fact) => fn context =>
   268       let
   269         val (fact', (decls, context')) =
   270           (fact, ([], context)) |-> fold_map (fn (ths, atts) => fn res1 =>
   271             (ths, res1) |-> fold_map (fn th => fn res2 =>
   272               let
   273                 val ((th', dyn'), res3) = fold eval (atts @ more_atts) ((th, NONE), res2);
   274                 val th_atts' =
   275                   (case dyn' of
   276                     NONE => (th', [])
   277                   | SOME (dyn_th', atts') => (dyn_th', rev atts'));
   278               in (th_atts', res3) end))
   279           |>> flat;
   280         val decls' = rev (map (apsnd rev) decls);
   281         val facts' =
   282           if eq_list (eq_fst strict_eq_thm) (decls', fact') then
   283             [((b, []), map2 (fn (th, atts1) => fn (_, atts2) => (th, atts1 @ atts2)) decls' fact')]
   284           else [(empty_binding, decls'), ((b, []), fact')];
   285       in (facts', context') end)
   286   |> fst |> flat |> map (apsnd (map (apfst single)));
   287 
   288 end;
   289 
   290 
   291 
   292 (** basic attributes **)
   293 
   294 (* internal *)
   295 
   296 fun internal att = Args.src (("Pure.attribute", [Token.mk_attribute att]), Position.none);
   297 
   298 
   299 (* rule composition *)
   300 
   301 val COMP_att =
   302   Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- thm
   303     >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => Drule.compose_single (A, i, B)));
   304 
   305 val THEN_att =
   306   Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- thm
   307     >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => A RSN (i, B)));
   308 
   309 val OF_att =
   310   thms >> (fn Bs => Thm.rule_attribute (fn _ => fn A => Bs MRS A));
   311 
   312 
   313 (* rename_abs *)
   314 
   315 val rename_abs =
   316   Scan.repeat (Args.maybe Args.name)
   317   >> (fn args => Thm.rule_attribute (K (Drule.rename_bvars' args)));
   318 
   319 
   320 (* unfold / fold definitions *)
   321 
   322 fun unfolded_syntax rule =
   323   thms >> (fn ths => Thm.rule_attribute (fn context => rule (Context.proof_of context) ths));
   324 
   325 val unfolded = unfolded_syntax Local_Defs.unfold;
   326 val folded = unfolded_syntax Local_Defs.fold;
   327 
   328 
   329 (* rule format *)
   330 
   331 val rule_format = Args.mode "no_asm"
   332   >> (fn true => Object_Logic.rule_format_no_asm | false => Object_Logic.rule_format);
   333 
   334 val elim_format = Thm.rule_attribute (K Tactic.make_elim);
   335 
   336 
   337 (* case names *)
   338 
   339 val case_names =
   340   Scan.repeat1 (Args.name --
   341     Scan.optional (Parse.$$$ "[" |-- Scan.repeat1 (Args.maybe Args.name) --| Parse.$$$ "]") []) >>
   342   (fn cs =>
   343     Rule_Cases.cases_hyp_names (map fst cs)
   344       (map (map (the_default Rule_Cases.case_hypsN) o snd) cs));
   345 
   346 
   347 (* misc rules *)
   348 
   349 val no_vars = Thm.rule_attribute (fn context => fn th =>
   350   let
   351     val ctxt = Variable.set_body false (Context.proof_of context);
   352     val ((_, [th']), _) = Variable.import true [th] ctxt;
   353   in th' end);
   354 
   355 val eta_long =
   356   Thm.rule_attribute (K (Conv.fconv_rule Drule.eta_long_conversion));
   357 
   358 val rotated = Scan.optional Parse.int 1 >> (fn n => Thm.rule_attribute (K (rotate_prems n)));
   359 
   360 
   361 (* theory setup *)
   362 
   363 val _ = Context.>> (Context.map_theory
   364  (setup (Binding.name "attribute") (Scan.lift Args.internal_attribute >> Morphism.form)
   365     "internal attribute" #>
   366   setup (Binding.name "tagged") (Scan.lift (Args.name -- Args.name) >> Thm.tag) "tagged theorem" #>
   367   setup (Binding.name "untagged") (Scan.lift Args.name >> Thm.untag) "untagged theorem" #>
   368   setup (Binding.name "kind") (Scan.lift Args.name >> Thm.kind) "theorem kind" #>
   369   setup (Binding.name "COMP") COMP_att "direct composition with rules (no lifting)" #>
   370   setup (Binding.name "THEN") THEN_att "resolution with rule" #>
   371   setup (Binding.name "OF") OF_att "rule applied to facts" #>
   372   setup (Binding.name "rename_abs") (Scan.lift rename_abs)
   373     "rename bound variables in abstractions" #>
   374   setup (Binding.name "unfolded") unfolded "unfolded definitions" #>
   375   setup (Binding.name "folded") folded "folded definitions" #>
   376   setup (Binding.name "consumes") (Scan.lift (Scan.optional Parse.nat 1) >> Rule_Cases.consumes)
   377     "number of consumed facts" #>
   378   setup (Binding.name "constraints") (Scan.lift Parse.nat >> Rule_Cases.constraints)
   379     "number of equality constraints" #>
   380   setup (Binding.name "case_names") (Scan.lift case_names) "named rule cases" #>
   381   setup (Binding.name "case_conclusion")
   382     (Scan.lift (Args.name -- Scan.repeat Args.name) >> Rule_Cases.case_conclusion)
   383     "named conclusion of rule cases" #>
   384   setup (Binding.name "params")
   385     (Scan.lift (Parse.and_list1 (Scan.repeat Args.name)) >> Rule_Cases.params)
   386     "named rule parameters" #>
   387   setup (Binding.name "standard") (Scan.succeed (Thm.rule_attribute (K Drule.export_without_context)))
   388     "result put into standard form (legacy)" #>
   389   setup (Binding.name "rule_format") rule_format "result put into canonical rule format" #>
   390   setup (Binding.name "elim_format") (Scan.succeed elim_format)
   391     "destruct rule turned into elimination rule format" #>
   392   setup (Binding.name "no_vars") (Scan.succeed no_vars) "frozen schematic vars" #>
   393   setup (Binding.name "eta_long") (Scan.succeed eta_long)
   394     "put theorem into eta long beta normal form" #>
   395   setup (Binding.name "atomize") (Scan.succeed Object_Logic.declare_atomize)
   396     "declaration of atomize rule" #>
   397   setup (Binding.name "rulify") (Scan.succeed Object_Logic.declare_rulify)
   398     "declaration of rulify rule" #>
   399   setup (Binding.name "rotated") (Scan.lift rotated) "rotated theorem premises" #>
   400   setup (Binding.name "defn") (add_del Local_Defs.defn_add Local_Defs.defn_del)
   401     "declaration of definitional transformations" #>
   402   setup (Binding.name "abs_def") (Scan.succeed (Thm.rule_attribute (K Drule.abs_def)))
   403     "abstract over free variables of a definition"));
   404 
   405 
   406 
   407 (** configuration options **)
   408 
   409 (* naming *)
   410 
   411 structure Configs = Theory_Data
   412 (
   413   type T = Config.raw Symtab.table;
   414   val empty = Symtab.empty;
   415   val extend = I;
   416   fun merge data = Symtab.merge (K true) data;
   417 );
   418 
   419 fun print_configs ctxt =
   420   let
   421     val thy = Proof_Context.theory_of ctxt;
   422     fun prt (name, config) =
   423       let val value = Config.get ctxt config in
   424         Pretty.block [Pretty.str (name ^ ": " ^ Config.print_type value ^ " ="), Pretty.brk 1,
   425           Pretty.str (Config.print_value value)]
   426       end;
   427     val configs = Name_Space.extern_table ctxt (#1 (Attributes.get thy), Configs.get thy);
   428   in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end;
   429 
   430 
   431 (* concrete syntax *)
   432 
   433 local
   434 
   435 val equals = Parse.$$$ "=";
   436 
   437 fun scan_value (Config.Bool _) =
   438       equals -- Args.$$$ "false" >> K (Config.Bool false) ||
   439       equals -- Args.$$$ "true" >> K (Config.Bool true) ||
   440       Scan.succeed (Config.Bool true)
   441   | scan_value (Config.Int _) = equals |-- Parse.int >> Config.Int
   442   | scan_value (Config.Real _) = equals |-- Parse.real >> Config.Real
   443   | scan_value (Config.String _) = equals |-- Args.name >> Config.String;
   444 
   445 fun scan_config thy config =
   446   let val config_type = Config.get_global thy config
   447   in scan_value config_type >> (K o Thm.declaration_attribute o K o Config.put_generic config) end;
   448 
   449 fun register binding config thy =
   450   let val name = Sign.full_name thy binding in
   451     thy
   452     |> setup binding (Scan.lift (scan_config thy config) >> Morphism.form) "configuration option"
   453     |> Configs.map (Symtab.update (name, config))
   454   end;
   455 
   456 fun declare make coerce binding default =
   457   let
   458     val name = Binding.name_of binding;
   459     val config_value = Config.declare_generic {global = false} name (make o default);
   460     val config = coerce config_value;
   461   in (config, register binding config_value) end;
   462 
   463 in
   464 
   465 val config_bool = declare Config.Bool Config.bool;
   466 val config_int = declare Config.Int Config.int;
   467 val config_real = declare Config.Real Config.real;
   468 val config_string = declare Config.String Config.string;
   469 
   470 fun register_config config = register (Binding.name (Config.name_of config)) config;
   471 
   472 end;
   473 
   474 
   475 (* implicit setup *)
   476 
   477 local
   478 
   479 fun setup_config declare_config binding default =
   480   let
   481     val (config, setup) = declare_config binding default;
   482     val _ = Context.>> (Context.map_theory setup);
   483   in config end;
   484 
   485 in
   486 
   487 val setup_config_bool = setup_config config_bool;
   488 val setup_config_int = setup_config config_int;
   489 val setup_config_string = setup_config config_string;
   490 val setup_config_real = setup_config config_real;
   491 
   492 end;
   493 
   494 
   495 (* theory setup *)
   496 
   497 val _ = Context.>> (Context.map_theory
   498  (register_config Ast.trace_raw #>
   499   register_config Ast.stats_raw #>
   500   register_config Syntax.positions_raw #>
   501   register_config Printer.show_brackets_raw #>
   502   register_config Printer.show_sorts_raw #>
   503   register_config Printer.show_types_raw #>
   504   register_config Printer.show_structs_raw #>
   505   register_config Printer.show_question_marks_raw #>
   506   register_config Syntax.ambiguity_level_raw #>
   507   register_config Syntax.ambiguity_warnings_raw #>
   508   register_config Syntax_Trans.eta_contract_raw #>
   509   register_config Name_Space.names_long_raw #>
   510   register_config Name_Space.names_short_raw #>
   511   register_config Name_Space.names_unique_raw #>
   512   register_config ML_Context.trace_raw #>
   513   register_config Proof_Context.show_abbrevs_raw #>
   514   register_config Goal_Display.goals_limit_raw #>
   515   register_config Goal_Display.show_main_goal_raw #>
   516   register_config Goal_Display.show_consts_raw #>
   517   register_config Display.show_hyps_raw #>
   518   register_config Display.show_tags_raw #>
   519   register_config Unify.trace_bound_raw #>
   520   register_config Unify.search_bound_raw #>
   521   register_config Unify.trace_simp_raw #>
   522   register_config Unify.trace_types_raw #>
   523   register_config Raw_Simplifier.simp_depth_limit_raw #>
   524   register_config Raw_Simplifier.simp_trace_depth_limit_raw #>
   525   register_config Raw_Simplifier.simp_debug_raw #>
   526   register_config Raw_Simplifier.simp_trace_raw));
   527 
   528 end;