src/Pure/Isar/attrib.ML
author wenzelm
Sun, 06 Nov 2011 21:51:46 +0100
changeset 46246 7fe19930dfc9
parent 44933 d7c7ec248ef0
child 46261 e29521ef9059
permissions -rw-r--r--
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
misc tuning;
     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 map_specs: ('a -> 'att) ->
    20     (('c * 'a list) * 'b) list -> (('c * 'att list) * 'b) list
    21   val map_facts: ('a -> 'att) ->
    22     (('c * 'a list) * ('d * 'a list) list) list ->
    23     (('c * 'att list) * ('d * 'att list) list) list
    24   val map_facts_refs: ('a -> 'att) -> ('b -> 'fact) ->
    25     (('c * 'a list) * ('b * 'a list) list) list ->
    26     (('c * 'att list) * ('fact * 'att list) list) list
    27   val eval_thms: Proof.context -> (Facts.ref * src list) list -> thm 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 config_bool: Binding.binding ->
    40     (Context.generic -> bool) -> bool Config.T * (theory -> theory)
    41   val config_int: Binding.binding ->
    42     (Context.generic -> int) -> int Config.T * (theory -> theory)
    43   val config_real: Binding.binding ->
    44     (Context.generic -> real) -> real Config.T * (theory -> theory)
    45   val config_string: Binding.binding ->
    46     (Context.generic -> string) -> string Config.T * (theory -> theory)
    47   val setup_config_bool: Binding.binding -> (Context.generic -> bool) -> bool Config.T
    48   val setup_config_int: Binding.binding -> (Context.generic -> int) -> int Config.T
    49   val setup_config_string: Binding.binding -> (Context.generic -> string) -> string Config.T
    50   val setup_config_real: Binding.binding -> (Context.generic -> real) -> real Config.T
    51 end;
    52 
    53 structure Attrib: ATTRIB =
    54 struct
    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 = Theory_Data
    70 (
    71   type T = ((src -> attribute) * string) Name_Space.table;
    72   val empty : T = Name_Space.empty_table "attribute";
    73   val extend = I;
    74   fun merge data : T = Name_Space.merge_tables data;
    75 );
    76 
    77 fun print_attributes thy =
    78   let
    79     val ctxt = Proof_Context.init_global thy;
    80     val attribs = Attributes.get thy;
    81     fun prt_attr (name, (_, "")) = Pretty.str name
    82       | prt_attr (name, (_, comment)) =
    83           Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
    84   in
    85     [Pretty.big_list "attributes:" (map prt_attr (Name_Space.extern_table ctxt attribs))]
    86     |> Pretty.chunks |> Pretty.writeln
    87   end;
    88 
    89 fun add_attribute name att comment thy = thy
    90   |> Attributes.map
    91     (Name_Space.define (Proof_Context.init_global thy) true (Sign.naming_of thy)
    92       (name, (att, comment)) #> snd);
    93 
    94 
    95 (* name space *)
    96 
    97 val intern = Name_Space.intern o #1 o Attributes.get;
    98 val intern_src = Args.map_name o intern;
    99 
   100 fun extern ctxt = Name_Space.extern ctxt (#1 (Attributes.get (Proof_Context.theory_of ctxt)));
   101 
   102 
   103 (* pretty printing *)
   104 
   105 fun pretty_attribs _ [] = []
   106   | pretty_attribs ctxt srcs =
   107       [Pretty.enum "," "[" "]" (map (Args.pretty_src ctxt o Args.map_name (extern ctxt)) srcs)];
   108 
   109 
   110 (* get attributes *)
   111 
   112 val defined = Symtab.defined o #2 o Attributes.get;
   113 
   114 fun attribute_i thy =
   115   let
   116     val (space, tab) = Attributes.get thy;
   117     fun attr src =
   118       let val ((name, _), pos) = Args.dest_src src in
   119         (case Symtab.lookup tab name of
   120           NONE => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
   121         | SOME (att, _) => (Position.report pos (Name_Space.markup space name); att src))
   122       end;
   123   in attr end;
   124 
   125 fun attribute thy = attribute_i thy o intern_src thy;
   126 
   127 
   128 (* attributed declarations *)
   129 
   130 fun map_specs f = map (apfst (apsnd (map f)));
   131 
   132 fun map_facts f = map (apfst (apsnd (map f)) o apsnd (map (apsnd (map f))));
   133 fun map_facts_refs f g = map_facts f #> map (apsnd (map (apfst g)));
   134 
   135 
   136 (* fact expressions *)
   137 
   138 fun eval_thms ctxt srcs = ctxt
   139   |> Proof_Context.note_thmss ""
   140     (map_facts_refs (attribute (Proof_Context.theory_of ctxt)) (Proof_Context.get_fact ctxt)
   141       [((Binding.empty, []), srcs)])
   142   |> fst |> maps snd;
   143 
   144 
   145 (* crude_closure *)
   146 
   147 (*Produce closure without knowing facts in advance! The following
   148   works reasonably well for attribute parsers that do not peek at the
   149   thm structure.*)
   150 
   151 fun crude_closure ctxt src =
   152  (try (fn () =>
   153     Thm.apply_attribute (attribute_i (Proof_Context.theory_of ctxt) src)
   154       (Context.Proof ctxt, Drule.asm_rl)) ();
   155   Args.closure src);
   156 
   157 
   158 (* attribute setup *)
   159 
   160 fun syntax scan = Args.syntax "attribute" scan;
   161 
   162 fun setup name scan =
   163   add_attribute name
   164     (fn src => fn (ctxt, th) => let val (a, ctxt') = syntax scan src ctxt in a (ctxt', th) end);
   165 
   166 fun attribute_setup name (txt, pos) cmt =
   167   Context.theory_map (ML_Context.expression pos
   168     "val (name, scan, comment): binding * attribute context_parser * string"
   169     "Context.map_theory (Attrib.setup name scan comment)"
   170     (ML_Lex.read Position.none ("(" ^ ML_Syntax.make_binding name ^ ", ") @
   171       ML_Lex.read pos txt @
   172       ML_Lex.read Position.none (", " ^ ML_Syntax.print_string cmt ^ ")")));
   173 
   174 
   175 (* add/del syntax *)
   176 
   177 fun add_del add del = Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add);
   178 
   179 
   180 
   181 (** parsing attributed theorems **)
   182 
   183 val thm_sel = Parse.$$$ "(" |-- Parse.list1
   184  (Parse.nat --| Parse.minus -- Parse.nat >> Facts.FromTo ||
   185   Parse.nat --| Parse.minus >> Facts.From ||
   186   Parse.nat >> Facts.Single) --| Parse.$$$ ")";
   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 (Global_Theory.get_fact context) Proof_Context.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     Parse.$$$ "[" |-- Args.attribs (intern thy) --| Parse.$$$ "]" >> (fn srcs =>
   200       let
   201         val atts = map (attribute_i thy) srcs;
   202         val (context', th') =
   203           Library.apply (map Thm.apply_attribute atts) (context, Drule.dummy_thm);
   204       in (context', pick "" [th']) end)
   205     ||
   206     (Scan.ahead Args.alt_name -- Args.named_fact get_fact
   207       >> (fn (s, fact) => ("", Facts.Fact s, fact)) ||
   208      Scan.ahead (Parse.position fact_name) :|-- (fn (name, pos) =>
   209       Args.named_fact (get_named pos) -- Scan.option thm_sel
   210         >> (fn (fact, sel) => (name, Facts.Named ((name, pos), sel), fact))))
   211     -- Args.opt_attribs (intern thy) >> (fn ((name, thmref, fact), srcs) =>
   212       let
   213         val ths = Facts.select thmref fact;
   214         val atts = map (attribute_i thy) srcs;
   215         val (context', ths') =
   216           Library.foldl_map (Library.apply (map Thm.apply_attribute atts)) (context, ths);
   217       in (context', pick name ths') end)
   218   end);
   219 
   220 in
   221 
   222 val thm = gen_thm Facts.the_single;
   223 val multi_thm = gen_thm (K I);
   224 val thms = Scan.repeat multi_thm >> flat;
   225 
   226 end;
   227 
   228 
   229 
   230 (** basic attributes **)
   231 
   232 (* internal *)
   233 
   234 fun internal att = Args.src (("Pure.attribute", [Token.mk_attribute att]), Position.none);
   235 
   236 
   237 (* rule composition *)
   238 
   239 val COMP_att =
   240   Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- thm
   241     >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => Drule.compose_single (A, i, B)));
   242 
   243 val THEN_att =
   244   Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- thm
   245     >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => A RSN (i, B)));
   246 
   247 val OF_att =
   248   thms >> (fn Bs => Thm.rule_attribute (fn _ => fn A => Bs MRS A));
   249 
   250 
   251 (* rename_abs *)
   252 
   253 val rename_abs =
   254   Scan.repeat (Args.maybe Args.name)
   255   >> (fn args => Thm.rule_attribute (K (Drule.rename_bvars' args)));
   256 
   257 
   258 (* unfold / fold definitions *)
   259 
   260 fun unfolded_syntax rule =
   261   thms >> (fn ths => Thm.rule_attribute (fn context => rule (Context.proof_of context) ths));
   262 
   263 val unfolded = unfolded_syntax Local_Defs.unfold;
   264 val folded = unfolded_syntax Local_Defs.fold;
   265 
   266 
   267 (* rule format *)
   268 
   269 val rule_format = Args.mode "no_asm"
   270   >> (fn true => Object_Logic.rule_format_no_asm | false => Object_Logic.rule_format);
   271 
   272 val elim_format = Thm.rule_attribute (K Tactic.make_elim);
   273 
   274 
   275 (* case names *)
   276 
   277 val case_names =
   278   Scan.repeat1 (Args.name --
   279     Scan.optional (Parse.$$$ "[" |-- Scan.repeat1 (Args.maybe Args.name) --| Parse.$$$ "]") []) >>
   280   (fn cs =>
   281     Rule_Cases.cases_hyp_names (map fst cs)
   282       (map (map (the_default Rule_Cases.case_hypsN) o snd) cs));
   283 
   284 
   285 (* misc rules *)
   286 
   287 val no_vars = Thm.rule_attribute (fn context => fn th =>
   288   let
   289     val ctxt = Variable.set_body false (Context.proof_of context);
   290     val ((_, [th']), _) = Variable.import true [th] ctxt;
   291   in th' end);
   292 
   293 val eta_long =
   294   Thm.rule_attribute (K (Conv.fconv_rule Drule.eta_long_conversion));
   295 
   296 val rotated = Scan.optional Parse.int 1 >> (fn n => Thm.rule_attribute (K (rotate_prems n)));
   297 
   298 
   299 (* theory setup *)
   300 
   301 val _ = Context.>> (Context.map_theory
   302  (setup (Binding.name "attribute") (Scan.lift Args.internal_attribute >> Morphism.form)
   303     "internal attribute" #>
   304   setup (Binding.name "tagged") (Scan.lift (Args.name -- Args.name) >> Thm.tag) "tagged theorem" #>
   305   setup (Binding.name "untagged") (Scan.lift Args.name >> Thm.untag) "untagged theorem" #>
   306   setup (Binding.name "kind") (Scan.lift Args.name >> Thm.kind) "theorem kind" #>
   307   setup (Binding.name "COMP") COMP_att "direct composition with rules (no lifting)" #>
   308   setup (Binding.name "THEN") THEN_att "resolution with rule" #>
   309   setup (Binding.name "OF") OF_att "rule applied to facts" #>
   310   setup (Binding.name "rename_abs") (Scan.lift rename_abs)
   311     "rename bound variables in abstractions" #>
   312   setup (Binding.name "unfolded") unfolded "unfolded definitions" #>
   313   setup (Binding.name "folded") folded "folded definitions" #>
   314   setup (Binding.name "consumes") (Scan.lift (Scan.optional Parse.nat 1) >> Rule_Cases.consumes)
   315     "number of consumed facts" #>
   316   setup (Binding.name "constraints") (Scan.lift Parse.nat >> Rule_Cases.constraints)
   317     "number of equality constraints" #>
   318   setup (Binding.name "case_names") (Scan.lift case_names) "named rule cases" #>
   319   setup (Binding.name "case_conclusion")
   320     (Scan.lift (Args.name -- Scan.repeat Args.name) >> Rule_Cases.case_conclusion)
   321     "named conclusion of rule cases" #>
   322   setup (Binding.name "params")
   323     (Scan.lift (Parse.and_list1 (Scan.repeat Args.name)) >> Rule_Cases.params)
   324     "named rule parameters" #>
   325   setup (Binding.name "standard") (Scan.succeed (Thm.rule_attribute (K Drule.export_without_context)))
   326     "result put into standard form (legacy)" #>
   327   setup (Binding.name "rule_format") rule_format "result put into canonical rule format" #>
   328   setup (Binding.name "elim_format") (Scan.succeed elim_format)
   329     "destruct rule turned into elimination rule format" #>
   330   setup (Binding.name "no_vars") (Scan.succeed no_vars) "frozen schematic vars" #>
   331   setup (Binding.name "eta_long") (Scan.succeed eta_long)
   332     "put theorem into eta long beta normal form" #>
   333   setup (Binding.name "atomize") (Scan.succeed Object_Logic.declare_atomize)
   334     "declaration of atomize rule" #>
   335   setup (Binding.name "rulify") (Scan.succeed Object_Logic.declare_rulify)
   336     "declaration of rulify rule" #>
   337   setup (Binding.name "rotated") (Scan.lift rotated) "rotated theorem premises" #>
   338   setup (Binding.name "defn") (add_del Local_Defs.defn_add Local_Defs.defn_del)
   339     "declaration of definitional transformations" #>
   340   setup (Binding.name "abs_def") (Scan.succeed (Thm.rule_attribute (K Drule.abs_def)))
   341     "abstract over free variables of a definition"));
   342 
   343 
   344 
   345 (** configuration options **)
   346 
   347 (* naming *)
   348 
   349 structure Configs = Theory_Data
   350 (
   351   type T = Config.raw Symtab.table;
   352   val empty = Symtab.empty;
   353   val extend = I;
   354   fun merge data = Symtab.merge (K true) data;
   355 );
   356 
   357 fun print_configs ctxt =
   358   let
   359     val thy = Proof_Context.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 = Name_Space.extern_table ctxt (#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 = Parse.$$$ "=";
   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 |-- Parse.int >> Config.Int
   380   | scan_value (Config.Real _) = equals |-- Parse.real >> Config.Real
   381   | scan_value (Config.String _) = equals |-- Args.name >> Config.String;
   382 
   383 fun scan_config thy config =
   384   let val config_type = Config.get_global thy config
   385   in scan_value config_type >> (K o Thm.declaration_attribute o K o Config.put_generic config) end;
   386 
   387 fun register binding config thy =
   388   let val name = Sign.full_name thy binding in
   389     thy
   390     |> setup binding (Scan.lift (scan_config thy config) >> Morphism.form) "configuration option"
   391     |> Configs.map (Symtab.update (name, config))
   392   end;
   393 
   394 fun declare make coerce binding default =
   395   let
   396     val name = Binding.name_of binding;
   397     val config_value = Config.declare_generic {global = false} name (make o default);
   398     val config = coerce config_value;
   399   in (config, register binding config_value) end;
   400 
   401 in
   402 
   403 val config_bool = declare Config.Bool Config.bool;
   404 val config_int = declare Config.Int Config.int;
   405 val config_real = declare Config.Real Config.real;
   406 val config_string = declare Config.String Config.string;
   407 
   408 fun register_config config = register (Binding.name (Config.name_of config)) config;
   409 
   410 end;
   411 
   412 
   413 (* implicit setup *)
   414 
   415 local
   416 
   417 fun setup_config declare_config binding default =
   418   let
   419     val (config, setup) = declare_config binding default;
   420     val _ = Context.>> (Context.map_theory setup);
   421   in config end;
   422 
   423 in
   424 
   425 val setup_config_bool = setup_config config_bool;
   426 val setup_config_int = setup_config config_int;
   427 val setup_config_string = setup_config config_string;
   428 val setup_config_real = setup_config config_real;
   429 
   430 end;
   431 
   432 
   433 (* theory setup *)
   434 
   435 val _ = Context.>> (Context.map_theory
   436  (register_config Ast.trace_raw #>
   437   register_config Ast.stats_raw #>
   438   register_config Syntax.positions_raw #>
   439   register_config Printer.show_brackets_raw #>
   440   register_config Printer.show_sorts_raw #>
   441   register_config Printer.show_types_raw #>
   442   register_config Printer.show_structs_raw #>
   443   register_config Printer.show_question_marks_raw #>
   444   register_config Syntax.ambiguity_level_raw #>
   445   register_config Syntax.ambiguity_warnings_raw #>
   446   register_config Syntax_Trans.eta_contract_raw #>
   447   register_config Name_Space.names_long_raw #>
   448   register_config Name_Space.names_short_raw #>
   449   register_config Name_Space.names_unique_raw #>
   450   register_config ML_Context.trace_raw #>
   451   register_config Proof_Context.show_abbrevs_raw #>
   452   register_config Goal_Display.goals_limit_raw #>
   453   register_config Goal_Display.show_main_goal_raw #>
   454   register_config Goal_Display.show_consts_raw #>
   455   register_config Display.show_hyps_raw #>
   456   register_config Display.show_tags_raw #>
   457   register_config Unify.trace_bound_raw #>
   458   register_config Unify.search_bound_raw #>
   459   register_config Unify.trace_simp_raw #>
   460   register_config Unify.trace_types_raw #>
   461   register_config Raw_Simplifier.simp_depth_limit_raw #>
   462   register_config Raw_Simplifier.simp_trace_depth_limit_raw #>
   463   register_config Raw_Simplifier.simp_debug_raw #>
   464   register_config Raw_Simplifier.simp_trace_raw));
   465 
   466 end;