wenzelm@5823: (* Title: Pure/Isar/attrib.ML wenzelm@5823: Author: Markus Wenzel, TU Muenchen wenzelm@5823: wenzelm@18734: Symbolic representation of attributes -- with name and syntax. wenzelm@5823: *) wenzelm@5823: wenzelm@5823: signature ATTRIB = wenzelm@5823: sig wenzelm@27729: type src = Args.src haftmann@29581: type binding = binding * src list haftmann@28965: val empty_binding: binding wenzelm@27729: val print_attributes: theory -> unit wenzelm@16458: val intern: theory -> xstring -> string wenzelm@16458: val intern_src: theory -> src -> src wenzelm@21031: val pretty_attribs: Proof.context -> src list -> Pretty.T list wenzelm@26891: val defined: theory -> string -> bool wenzelm@18734: val attribute: theory -> src -> attribute wenzelm@18734: val attribute_i: theory -> src -> attribute wenzelm@26336: val eval_thms: Proof.context -> (Facts.ref * src list) list -> thm list wenzelm@18905: val map_specs: ('a -> 'att) -> wenzelm@18905: (('c * 'a list) * 'd) list -> (('c * 'att list) * 'd) list wenzelm@18905: val map_facts: ('a -> 'att) -> wenzelm@17105: (('c * 'a list) * ('d * 'a list) list) list -> wenzelm@18905: (('c * 'att list) * ('d * 'att list) list) list wenzelm@20289: val crude_closure: Proof.context -> src -> src wenzelm@18734: val add_attributes: (bstring * (src -> attribute) * string) list -> theory -> theory wenzelm@25983: val syntax: (Context.generic * Args.T list -> wenzelm@25983: attribute * (Context.generic * Args.T list)) -> src -> attribute wenzelm@25983: val no_args: attribute -> src -> attribute wenzelm@25983: val add_del_args: attribute -> attribute -> src -> attribute wenzelm@27812: val thm_sel: Args.T list -> Facts.interval list * Args.T list wenzelm@25983: val thm: Context.generic * Args.T list -> thm * (Context.generic * Args.T list) wenzelm@25983: val thms: Context.generic * Args.T list -> thm list * (Context.generic * Args.T list) wenzelm@25983: val multi_thm: Context.generic * Args.T list -> thm list * (Context.generic * Args.T list) wenzelm@24110: val print_configs: Proof.context -> unit wenzelm@25983: val internal: (morphism -> attribute) -> src wenzelm@24126: val register_config: Config.value Config.T -> theory -> theory wenzelm@24110: val config_bool: bstring -> bool -> bool Config.T * (theory -> theory) wenzelm@24110: val config_int: bstring -> int -> int Config.T * (theory -> theory) wenzelm@24110: val config_string: bstring -> string -> string Config.T * (theory -> theory) wenzelm@24110: val config_bool_global: bstring -> bool -> bool Config.T * (theory -> theory) wenzelm@24110: val config_int_global: bstring -> int -> int Config.T * (theory -> theory) wenzelm@24110: val config_string_global: bstring -> string -> string Config.T * (theory -> theory) wenzelm@5823: end; wenzelm@5823: wenzelm@5823: structure Attrib: ATTRIB = wenzelm@5823: struct wenzelm@5823: wenzelm@27812: structure T = OuterLex; wenzelm@27812: structure P = OuterParse; wenzelm@27812: wenzelm@28084: wenzelm@28084: (* source and bindings *) wenzelm@28084: wenzelm@15703: type src = Args.src; wenzelm@15703: haftmann@29581: type binding = binding * src list; haftmann@28965: val empty_binding: binding = (Binding.empty, []); wenzelm@28084: wenzelm@28084: wenzelm@27729: wenzelm@18734: (** named attributes **) wenzelm@18636: wenzelm@18734: (* theory data *) wenzelm@5823: wenzelm@24110: structure Attributes = TheoryDataFun wenzelm@22846: ( wenzelm@18734: type T = (((src -> attribute) * string) * stamp) NameSpace.table; wenzelm@16344: val empty = NameSpace.empty_table; wenzelm@6546: val copy = I; wenzelm@16458: val extend = I; wenzelm@23655: fun merge _ tables : T = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUP dup => wenzelm@23655: error ("Attempt to merge different versions of attribute " ^ quote dup); wenzelm@22846: ); wenzelm@5823: wenzelm@22846: fun print_attributes thy = wenzelm@22846: let wenzelm@24110: val attribs = Attributes.get thy; wenzelm@22846: fun prt_attr (name, ((_, comment), _)) = Pretty.block wenzelm@22846: [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; wenzelm@22846: in wenzelm@22846: [Pretty.big_list "attributes:" (map prt_attr (NameSpace.extern_table attribs))] wenzelm@22846: |> Pretty.chunks |> Pretty.writeln wenzelm@22846: end; wenzelm@7611: wenzelm@5823: wenzelm@21031: (* name space *) wenzelm@15703: wenzelm@24110: val intern = NameSpace.intern o #1 o Attributes.get; wenzelm@15703: val intern_src = Args.map_name o intern; wenzelm@15703: wenzelm@24110: val extern = NameSpace.extern o #1 o Attributes.get o ProofContext.theory_of; wenzelm@21031: wenzelm@21031: wenzelm@21031: (* pretty printing *) wenzelm@21031: wenzelm@21031: fun pretty_attribs _ [] = [] wenzelm@21031: | pretty_attribs ctxt srcs = wenzelm@21031: [Pretty.enclose "[" "]" wenzelm@21031: (Pretty.commas (map (Args.pretty_src ctxt o Args.map_name (extern ctxt)) srcs))]; wenzelm@21031: wenzelm@15703: wenzelm@18734: (* get attributes *) wenzelm@5823: wenzelm@26891: val defined = Symtab.defined o #2 o Attributes.get; wenzelm@26891: wenzelm@18734: fun attribute_i thy = wenzelm@5823: let wenzelm@24110: val attrs = #2 (Attributes.get thy); wenzelm@5879: fun attr src = wenzelm@16344: let val ((name, _), pos) = Args.dest_src src in wenzelm@17412: (case Symtab.lookup attrs name of skalberg@15531: NONE => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos) wenzelm@27751: | SOME ((att, _), _) => (Position.report (Markup.attribute name) pos; att src)) wenzelm@5823: end; wenzelm@5823: in attr end; wenzelm@5823: wenzelm@18734: fun attribute thy = attribute_i thy o intern_src thy; wenzelm@18636: wenzelm@24723: fun eval_thms ctxt args = ProofContext.note_thmss Thm.theoremK haftmann@28965: [((Binding.empty, []), wenzelm@28083: map (apsnd (map (attribute (ProofContext.theory_of ctxt)))) args)] ctxt wenzelm@26685: |> fst |> maps snd; wenzelm@24723: wenzelm@5823: wenzelm@17105: (* attributed declarations *) wenzelm@17105: wenzelm@17105: fun map_specs f = map (apfst (apsnd (map f))); wenzelm@17105: fun map_facts f = map (apfst (apsnd (map f)) o apsnd (map (apsnd (map f)))); wenzelm@17105: wenzelm@17105: wenzelm@15703: (* crude_closure *) wenzelm@15703: wenzelm@15703: (*Produce closure without knowing facts in advance! The following wenzelm@18734: works reasonably well for attribute parsers that do not peek at the wenzelm@18734: thm structure.*) wenzelm@15703: wenzelm@15703: fun crude_closure ctxt src = wenzelm@18734: (try (fn () => attribute_i (ProofContext.theory_of ctxt) src wenzelm@18734: (Context.Proof ctxt, Drule.asm_rl)) (); wenzelm@15703: Args.closure src); wenzelm@15703: wenzelm@15703: wenzelm@5823: (* add_attributes *) wenzelm@5823: wenzelm@5823: fun add_attributes raw_attrs thy = wenzelm@5823: let wenzelm@18734: val new_attrs = haftmann@29004: raw_attrs |> map (fn (name, att, comment) => (Binding.name name, ((att, comment), stamp ()))); haftmann@29004: fun add attrs = fold (snd oo NameSpace.bind (Sign.naming_of thy)) new_attrs attrs wenzelm@23655: handle Symtab.DUP dup => error ("Duplicate declaration of attributes " ^ quote dup); wenzelm@24110: in Attributes.map add thy end; wenzelm@5823: wenzelm@5879: wenzelm@25983: (* attribute syntax *) wenzelm@5879: wenzelm@25983: fun syntax scan src (st, th) = wenzelm@25983: let val (f: attribute, st') = Args.syntax "attribute" scan src st wenzelm@25983: in f (st', th) end; wenzelm@5823: wenzelm@25983: fun no_args x = syntax (Scan.succeed x); wenzelm@25983: wenzelm@25983: fun add_del_args add del = syntax wenzelm@25983: (Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add)); wenzelm@5879: wenzelm@5879: wenzelm@25983: wenzelm@25983: (** parsing attributed theorems **) wenzelm@5879: wenzelm@27812: val thm_sel = P.$$$ "(" |-- P.list1 wenzelm@27812: (P.nat --| P.minus -- P.nat >> Facts.FromTo || wenzelm@27812: P.nat --| P.minus >> Facts.From || wenzelm@27812: P.nat >> Facts.Single) --| P.$$$ ")"; wenzelm@27812: wenzelm@18636: local wenzelm@18636: wenzelm@21698: val fact_name = Args.internal_fact >> K "" || Args.name; wenzelm@21698: wenzelm@26336: fun gen_thm pick = Scan.depend (fn context => wenzelm@26336: let wenzelm@26336: val thy = Context.theory_of context; wenzelm@26685: val get = Context.cases (PureThy.get_fact context) ProofContext.get_fact context; wenzelm@26361: val get_fact = get o Facts.Fact; wenzelm@27820: fun get_named pos name = get (Facts.Named ((name, pos), NONE)); wenzelm@26336: in wenzelm@27820: P.$$$ "[" |-- Args.attribs (intern thy) --| P.$$$ "]" >> (fn srcs => wenzelm@24008: let wenzelm@26336: val atts = map (attribute_i thy) srcs; wenzelm@26336: val (context', th') = Library.apply atts (context, Drule.dummy_thm); wenzelm@26336: in (context', pick "" [th']) end) wenzelm@26336: || wenzelm@26336: (Scan.ahead Args.alt_name -- Args.named_fact get_fact wenzelm@27820: >> (fn (s, fact) => ("", Facts.Fact s, fact)) || wenzelm@27820: Scan.ahead (P.position fact_name) :|-- (fn (name, pos) => wenzelm@27820: Args.named_fact (get_named pos) -- Scan.option thm_sel wenzelm@27820: >> (fn (fact, sel) => (name, Facts.Named ((name, pos), sel), fact)))) wenzelm@26336: -- Args.opt_attribs (intern thy) >> (fn ((name, thmref, fact), srcs) => wenzelm@26336: let wenzelm@26336: val ths = Facts.select thmref fact; wenzelm@26336: val atts = map (attribute_i thy) srcs; wenzelm@30190: val (context', ths') = Library.foldl_map (Library.apply atts) (context, ths); wenzelm@26336: in (context', pick name ths') end) wenzelm@26336: end); berghofe@15456: wenzelm@18636: in wenzelm@18636: wenzelm@26336: val thm = gen_thm Facts.the_single; wenzelm@18998: val multi_thm = gen_thm (K I); wenzelm@19482: val thms = Scan.repeat multi_thm >> flat; wenzelm@18636: wenzelm@18636: end; wenzelm@18636: wenzelm@5823: wenzelm@5879: wenzelm@25983: (** basic attributes **) wenzelm@25983: wenzelm@25983: (* internal *) wenzelm@25983: wenzelm@27812: fun internal att = Args.src (("Pure.attribute", [T.mk_attribute att]), Position.none); wenzelm@25983: wenzelm@25983: val internal_att = wenzelm@25983: syntax (Scan.lift Args.internal_attribute >> Morphism.form); wenzelm@25983: wenzelm@25983: wenzelm@25983: (* tags *) wenzelm@25983: wenzelm@27865: val tagged = syntax (Scan.lift (Args.name -- Args.name) >> Thm.tag); wenzelm@27865: val untagged = syntax (Scan.lift Args.name >> Thm.untag); wenzelm@25983: wenzelm@27865: val kind = syntax (Scan.lift Args.name >> Thm.kind); wenzelm@25983: wenzelm@25983: wenzelm@25983: (* rule composition *) wenzelm@25983: wenzelm@25983: val COMP_att = wenzelm@27812: syntax (Scan.lift (Scan.optional (Args.bracks P.nat) 1) -- thm wenzelm@25983: >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => Drule.compose_single (A, i, B)))); wenzelm@25983: wenzelm@25983: val THEN_att = wenzelm@27812: syntax (Scan.lift (Scan.optional (Args.bracks P.nat) 1) -- thm wenzelm@25983: >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => A RSN (i, B)))); wenzelm@25983: wenzelm@25983: val OF_att = wenzelm@25983: syntax (thms >> (fn Bs => Thm.rule_attribute (fn _ => fn A => Bs MRS A))); wenzelm@25983: wenzelm@25983: wenzelm@25983: (* rename_abs *) wenzelm@25983: wenzelm@25983: val rename_abs = syntax wenzelm@25983: (Scan.lift (Scan.repeat (Args.maybe Args.name) >> (apsnd o Drule.rename_bvars'))); wenzelm@25983: wenzelm@25983: wenzelm@25983: (* unfold / fold definitions *) wenzelm@25983: wenzelm@25983: fun unfolded_syntax rule = wenzelm@25983: syntax (thms >> wenzelm@25983: (fn ths => Thm.rule_attribute (fn context => rule (Context.proof_of context) ths))); wenzelm@25983: wenzelm@25983: val unfolded = unfolded_syntax LocalDefs.unfold; wenzelm@25983: val folded = unfolded_syntax LocalDefs.fold; wenzelm@25983: wenzelm@25983: wenzelm@25983: (* rule cases *) wenzelm@5823: wenzelm@27812: val consumes = syntax (Scan.lift (Scan.optional P.nat 1) >> RuleCases.consumes); wenzelm@25983: val case_names = syntax (Scan.lift (Scan.repeat1 Args.name) >> RuleCases.case_names); wenzelm@25983: val case_conclusion = wenzelm@25983: syntax (Scan.lift (Args.name -- Scan.repeat Args.name) >> RuleCases.case_conclusion); wenzelm@27812: val params = syntax (Scan.lift (P.and_list1 (Scan.repeat Args.name)) >> RuleCases.params); wenzelm@25983: wenzelm@25983: wenzelm@25983: (* rule format *) wenzelm@25983: wenzelm@25983: val rule_format = syntax (Args.mode "no_asm" wenzelm@25983: >> (fn true => ObjectLogic.rule_format_no_asm | false => ObjectLogic.rule_format)); wenzelm@25983: wenzelm@25983: val elim_format = no_args (Thm.rule_attribute (K Tactic.make_elim)); wenzelm@25983: wenzelm@25983: wenzelm@25983: (* misc rules *) wenzelm@25983: wenzelm@25983: val standard = no_args (Thm.rule_attribute (K Drule.standard)); wenzelm@25983: wenzelm@26715: val no_vars = no_args (Thm.rule_attribute (fn context => fn th => wenzelm@26715: let wenzelm@26715: val ctxt = Variable.set_body false (Context.proof_of context); wenzelm@26715: val ((_, [th']), _) = Variable.import_thms true [th] ctxt; wenzelm@25983: in th' end)); wenzelm@25983: wenzelm@25983: val eta_long = wenzelm@25983: no_args (Thm.rule_attribute (K (Conv.fconv_rule Drule.eta_long_conversion))); wenzelm@25983: wenzelm@25983: val rotated = syntax wenzelm@27812: (Scan.lift (Scan.optional P.int 1) >> (fn n => Thm.rule_attribute (K (rotate_prems n)))); wenzelm@5879: berghofe@29690: val abs_def = no_args (Thm.rule_attribute (K Drule.abs_def)); berghofe@29690: wenzelm@25983: wenzelm@25983: (* theory setup *) wenzelm@5823: wenzelm@26463: val _ = Context.>> (Context.map_theory wenzelm@25983: (add_attributes wenzelm@25983: [("attribute", internal_att, "internal attribute"), wenzelm@25983: ("tagged", tagged, "tagged theorem"), wenzelm@25983: ("untagged", untagged, "untagged theorem"), wenzelm@25983: ("kind", kind, "theorem kind"), wenzelm@25983: ("COMP", COMP_att, "direct composition with rules (no lifting)"), wenzelm@25983: ("THEN", THEN_att, "resolution with rule"), wenzelm@25983: ("OF", OF_att, "rule applied to facts"), wenzelm@25983: ("rename_abs", rename_abs, "rename bound variables in abstractions"), wenzelm@25983: ("unfolded", unfolded, "unfolded definitions"), wenzelm@25983: ("folded", folded, "folded definitions"), wenzelm@25983: ("standard", standard, "result put into standard form"), wenzelm@25983: ("elim_format", elim_format, "destruct rule turned into elimination rule format"), wenzelm@25983: ("no_vars", no_vars, "frozen schematic vars"), wenzelm@25983: ("eta_long", eta_long, "put theorem into eta long beta normal form"), wenzelm@25983: ("consumes", consumes, "number of consumed facts"), wenzelm@25983: ("case_names", case_names, "named rule cases"), wenzelm@25983: ("case_conclusion", case_conclusion, "named conclusion of rule cases"), wenzelm@25983: ("params", params, "named rule parameters"), wenzelm@25983: ("atomize", no_args ObjectLogic.declare_atomize, "declaration of atomize rule"), wenzelm@25983: ("rulify", no_args ObjectLogic.declare_rulify, "declaration of rulify rule"), wenzelm@25983: ("rule_format", rule_format, "result put into standard rule format"), wenzelm@25983: ("rotated", rotated, "rotated theorem premises"), wenzelm@25983: ("defn", add_del_args LocalDefs.defn_add LocalDefs.defn_del, berghofe@29690: "declaration of definitional transformations"), berghofe@29690: ("abs_def", abs_def, "abstract over free variables of a definition")])); wenzelm@8633: wenzelm@5823: wenzelm@5823: wenzelm@24110: (** configuration options **) wenzelm@24110: wenzelm@24110: (* naming *) wenzelm@24110: wenzelm@24110: structure Configs = TheoryDataFun wenzelm@24713: ( wenzelm@24110: type T = Config.value Config.T Symtab.table; wenzelm@24110: val empty = Symtab.empty; wenzelm@24110: val copy = I; wenzelm@24110: val extend = I; wenzelm@24110: fun merge _ = Symtab.merge (K true); wenzelm@24713: ); wenzelm@5823: wenzelm@24110: fun print_configs ctxt = wenzelm@24110: let wenzelm@24110: val thy = ProofContext.theory_of ctxt; wenzelm@24110: fun prt (name, config) = wenzelm@24110: let val value = Config.get ctxt config in wenzelm@24110: Pretty.block [Pretty.str (name ^ ": " ^ Config.print_type value ^ " ="), Pretty.brk 1, wenzelm@24110: Pretty.str (Config.print_value value)] wenzelm@24110: end; wenzelm@24110: val configs = NameSpace.extern_table (#1 (Attributes.get thy), Configs.get thy); wenzelm@24110: in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end; wenzelm@24110: wenzelm@24110: wenzelm@24110: (* concrete syntax *) wenzelm@23988: wenzelm@24003: local wenzelm@24003: wenzelm@27820: val equals = P.$$$ "="; wenzelm@24003: wenzelm@24110: fun scan_value (Config.Bool _) = wenzelm@24110: equals -- Args.$$$ "false" >> K (Config.Bool false) || wenzelm@24110: equals -- Args.$$$ "true" >> K (Config.Bool true) || wenzelm@24110: Scan.succeed (Config.Bool true) wenzelm@27812: | scan_value (Config.Int _) = equals |-- P.int >> Config.Int wenzelm@24110: | scan_value (Config.String _) = equals |-- Args.name >> Config.String; wenzelm@24003: wenzelm@24110: fun scan_config thy config = wenzelm@24110: let val config_type = Config.get_thy thy config wenzelm@24110: in scan_value config_type >> (K o Thm.declaration_attribute o K o Config.put_generic config) end; wenzelm@24003: wenzelm@24003: in wenzelm@24003: wenzelm@24126: fun register_config config thy = wenzelm@24126: let wenzelm@24126: val bname = Config.name_of config; haftmann@28965: val name = Sign.full_bname thy bname; wenzelm@24126: in wenzelm@24126: thy wenzelm@24126: |> add_attributes [(bname, syntax (Scan.lift (scan_config thy config) >> Morphism.form), wenzelm@24126: "configuration option")] wenzelm@24126: |> Configs.map (Symtab.update (name, config)) wenzelm@24126: end; wenzelm@24110: wenzelm@24110: fun declare_config make coerce global name default = wenzelm@24110: let wenzelm@24110: val config_value = Config.declare global name (make default); wenzelm@24110: val config = coerce config_value; wenzelm@24126: in (config, register_config config_value) end; wenzelm@24110: wenzelm@24110: val config_bool = declare_config Config.Bool Config.bool false; wenzelm@24110: val config_int = declare_config Config.Int Config.int false; wenzelm@24110: val config_string = declare_config Config.String Config.string false; wenzelm@24110: wenzelm@24110: val config_bool_global = declare_config Config.Bool Config.bool true; wenzelm@24110: val config_int_global = declare_config Config.Int Config.int true; wenzelm@24110: val config_string_global = declare_config Config.String Config.string true; wenzelm@24003: wenzelm@24003: end; wenzelm@23988: wenzelm@23988: wenzelm@18636: (* theory setup *) wenzelm@5823: wenzelm@26463: val _ = Context.>> (Context.map_theory wenzelm@24178: (register_config Unify.trace_bound_value #> wenzelm@24178: register_config Unify.search_bound_value #> wenzelm@24178: register_config Unify.trace_simp_value #> wenzelm@24178: register_config Unify.trace_types_value #> wenzelm@26463: register_config MetaSimplifier.simp_depth_limit_value)); wenzelm@5823: wenzelm@5823: end;