| author | wenzelm | 
| Tue, 25 Mar 2014 14:52:35 +0100 | |
| changeset 56276 | 9e2d5e3debd3 | 
| parent 56265 | 785569927666 | 
| child 56278 | 2576d3a40ed6 | 
| permissions | -rw-r--r-- | 
| 5823 | 1 | (* Title: Pure/Isar/attrib.ML | 
| 2 | Author: Markus Wenzel, TU Muenchen | |
| 3 | ||
| 18734 | 4 | Symbolic representation of attributes -- with name and syntax. | 
| 5823 | 5 | *) | 
| 6 | ||
| 7 | signature ATTRIB = | |
| 8 | sig | |
| 27729 | 9 | type src = Args.src | 
| 29581 | 10 | type binding = binding * src list | 
| 28965 | 11 | val empty_binding: binding | 
| 45584 
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
 wenzelm parents: 
45537diff
changeset | 12 | val is_empty_binding: binding -> bool | 
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 13 | val print_attributes: Proof.context -> unit | 
| 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 14 | val check_name_generic: Context.generic -> xstring * Position.T -> string | 
| 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 15 | val check_name: Proof.context -> xstring * Position.T -> string | 
| 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 16 | val check_src: Proof.context -> src -> src | 
| 21031 | 17 | val pretty_attribs: Proof.context -> src list -> Pretty.T list | 
| 47815 | 18 | val attribute: Proof.context -> src -> attribute | 
| 19 | val attribute_global: theory -> src -> attribute | |
| 20 | val attribute_cmd: Proof.context -> src -> attribute | |
| 21 | val attribute_cmd_global: theory -> src -> attribute | |
| 45390 
e29521ef9059
tuned signature -- avoid spurious Thm.mixed_attribute;
 wenzelm parents: 
45375diff
changeset | 22 |   val map_specs: ('a list -> 'att list) ->
 | 
| 30759 | 23 |     (('c * 'a list) * 'b) list -> (('c * 'att list) * 'b) list
 | 
| 45390 
e29521ef9059
tuned signature -- avoid spurious Thm.mixed_attribute;
 wenzelm parents: 
45375diff
changeset | 24 |   val map_facts: ('a list -> 'att list) ->
 | 
| 17105 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 wenzelm parents: 
16934diff
changeset | 25 |     (('c * 'a list) * ('d * 'a list) list) list ->
 | 
| 18905 | 26 |     (('c * 'att list) * ('d * 'att list) list) list
 | 
| 45390 
e29521ef9059
tuned signature -- avoid spurious Thm.mixed_attribute;
 wenzelm parents: 
45375diff
changeset | 27 |   val map_facts_refs: ('a list -> 'att list) -> ('b -> 'fact) ->
 | 
| 30759 | 28 |     (('c * 'a list) * ('b * 'a list) list) list ->
 | 
| 29 |     (('c * 'att list) * ('fact * 'att list) list) list
 | |
| 47249 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 30 | val global_notes: string -> (binding * (thm list * src list) list) list -> | 
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 31 | theory -> (string * thm list) list * theory | 
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 32 | val local_notes: string -> (binding * (thm list * src list) list) list -> | 
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 33 | Proof.context -> (string * thm list) list * Proof.context | 
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 34 | val generic_notes: string -> (binding * (thm list * src list) list) list -> | 
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 35 | Context.generic -> (string * thm list) list * Context.generic | 
| 38330 | 36 | val eval_thms: Proof.context -> (Facts.ref * src list) list -> thm list | 
| 30525 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 37 | val setup: Binding.binding -> attribute context_parser -> string -> theory -> theory | 
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55740diff
changeset | 38 | val attribute_setup: bstring * Position.T -> Symbol_Pos.source -> string -> theory -> theory | 
| 55140 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 39 | val internal: (morphism -> attribute) -> src | 
| 30525 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 40 | val add_del: attribute -> attribute -> attribute context_parser | 
| 30513 | 41 | val thm_sel: Facts.interval list parser | 
| 42 | val thm: thm context_parser | |
| 43 | val thms: thm list context_parser | |
| 44 | val multi_thm: thm list context_parser | |
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 45 | val partial_evaluation: Proof.context -> | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 46 | (binding * (thm list * Args.src list) list) list -> | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 47 | (binding * (thm list * Args.src list) list) list | 
| 52060 | 48 | val print_options: Proof.context -> unit | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 49 | val config_bool: Binding.binding -> | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 50 | (Context.generic -> bool) -> bool Config.T * (theory -> theory) | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 51 | val config_int: Binding.binding -> | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 52 | (Context.generic -> int) -> int Config.T * (theory -> theory) | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 53 | val config_real: Binding.binding -> | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 54 | (Context.generic -> real) -> real Config.T * (theory -> theory) | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 55 | val config_string: Binding.binding -> | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 56 | (Context.generic -> string) -> string Config.T * (theory -> theory) | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 57 | val setup_config_bool: Binding.binding -> (Context.generic -> bool) -> bool Config.T | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 58 | val setup_config_int: Binding.binding -> (Context.generic -> int) -> int Config.T | 
| 52040 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 59 | val setup_config_real: Binding.binding -> (Context.generic -> real) -> real Config.T | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 60 | val setup_config_string: Binding.binding -> (Context.generic -> string) -> string Config.T | 
| 52040 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 61 | val option_bool: string -> bool Config.T * (theory -> theory) | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 62 | val option_int: string -> int Config.T * (theory -> theory) | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 63 | val option_real: string -> real Config.T * (theory -> theory) | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 64 | val option_string: string -> string Config.T * (theory -> theory) | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 65 | val setup_option_bool: string -> bool Config.T | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 66 | val setup_option_int: string -> int Config.T | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 67 | val setup_option_real: string -> real Config.T | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 68 | val setup_option_string: string -> string Config.T | 
| 5823 | 69 | end; | 
| 70 | ||
| 71 | structure Attrib: ATTRIB = | |
| 72 | struct | |
| 73 | ||
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 74 | (* source and bindings *) | 
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 75 | |
| 15703 | 76 | type src = Args.src; | 
| 77 | ||
| 29581 | 78 | type binding = binding * src list; | 
| 45584 
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
 wenzelm parents: 
45537diff
changeset | 79 | |
| 28965 | 80 | val empty_binding: binding = (Binding.empty, []); | 
| 45584 
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
 wenzelm parents: 
45537diff
changeset | 81 | fun is_empty_binding ((b, srcs): binding) = Binding.is_empty b andalso null srcs; | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 82 | |
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 83 | |
| 27729 | 84 | |
| 18734 | 85 | (** named attributes **) | 
| 18636 | 86 | |
| 18734 | 87 | (* theory data *) | 
| 5823 | 88 | |
| 33522 | 89 | structure Attributes = Theory_Data | 
| 22846 | 90 | ( | 
| 33095 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 wenzelm parents: 
33092diff
changeset | 91 | type T = ((src -> attribute) * string) Name_Space.table; | 
| 33159 | 92 | val empty : T = Name_Space.empty_table "attribute"; | 
| 16458 | 93 | val extend = I; | 
| 33522 | 94 | fun merge data : T = Name_Space.merge_tables data; | 
| 22846 | 95 | ); | 
| 5823 | 96 | |
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 97 | val get_attributes = Attributes.get o Context.theory_of; | 
| 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 98 | |
| 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 99 | fun print_attributes ctxt = | 
| 22846 | 100 | let | 
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 101 | val attribs = get_attributes (Context.Proof ctxt); | 
| 50301 | 102 | fun prt_attr (name, (_, "")) = Pretty.mark_str name | 
| 42813 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 wenzelm parents: 
42808diff
changeset | 103 | | prt_attr (name, (_, comment)) = | 
| 50301 | 104 | Pretty.block | 
| 105 | (Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment); | |
| 22846 | 106 | in | 
| 56052 | 107 | [Pretty.big_list "attributes:" (map prt_attr (Name_Space.markup_table ctxt attribs))] | 
| 22846 | 108 | |> Pretty.chunks |> Pretty.writeln | 
| 109 | end; | |
| 7611 | 110 | |
| 56027 
25889f5c39a8
prefer Name_Space.pretty with its builtin markup;
 wenzelm parents: 
56026diff
changeset | 111 | val attribute_space = Name_Space.space_of_table o get_attributes o Context.Proof; | 
| 
25889f5c39a8
prefer Name_Space.pretty with its builtin markup;
 wenzelm parents: 
56026diff
changeset | 112 | |
| 33092 
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 wenzelm parents: 
31794diff
changeset | 113 | fun add_attribute name att comment thy = thy | 
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46906diff
changeset | 114 | |> Attributes.map (Name_Space.define (Context.Theory thy) true (name, (att, comment)) #> snd); | 
| 31306 
a74ee84288a0
eliminated old Attrib.add_attributes (and Attrib.syntax);
 wenzelm parents: 
31305diff
changeset | 115 | |
| 5823 | 116 | |
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 117 | (* check *) | 
| 15703 | 118 | |
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 119 | fun check_name_generic context = #1 o Name_Space.check context (get_attributes context); | 
| 56032 
b034b9f0fa2a
clarified Args.check_src: retain name space information for error output;
 wenzelm parents: 
56029diff
changeset | 120 | val check_name = check_name_generic o Context.Proof; | 
| 15703 | 121 | |
| 56033 | 122 | fun check_src ctxt src = | 
| 123 | (Context_Position.report ctxt (Args.range_of_src src) Markup.language_attribute; | |
| 124 | #1 (Args.check_src ctxt (get_attributes (Context.Proof ctxt)) src)); | |
| 21031 | 125 | |
| 126 | ||
| 127 | (* pretty printing *) | |
| 128 | ||
| 129 | fun pretty_attribs _ [] = [] | |
| 56032 
b034b9f0fa2a
clarified Args.check_src: retain name space information for error output;
 wenzelm parents: 
56029diff
changeset | 130 | | pretty_attribs ctxt srcs = [Pretty.enum "," "[" "]" (map (Args.pretty_src ctxt) srcs)]; | 
| 21031 | 131 | |
| 15703 | 132 | |
| 18734 | 133 | (* get attributes *) | 
| 5823 | 134 | |
| 47816 
cd0dfb06b0c8
prefer Context_Position.report_generic, which observes is_visible flag and thus reduces number of echos;
 wenzelm parents: 
47815diff
changeset | 135 | fun attribute_generic context = | 
| 56029 
8bedca4bd5a3
clarified Args.src: more abstract type, position refers to name only;
 wenzelm parents: 
56027diff
changeset | 136 | let val table = get_attributes context | 
| 
8bedca4bd5a3
clarified Args.src: more abstract type, position refers to name only;
 wenzelm parents: 
56027diff
changeset | 137 | in fn src => #1 (Name_Space.get table (#1 (Args.name_of_src src))) src end; | 
| 5823 | 138 | |
| 47816 
cd0dfb06b0c8
prefer Context_Position.report_generic, which observes is_visible flag and thus reduces number of echos;
 wenzelm parents: 
47815diff
changeset | 139 | val attribute = attribute_generic o Context.Proof; | 
| 
cd0dfb06b0c8
prefer Context_Position.report_generic, which observes is_visible flag and thus reduces number of echos;
 wenzelm parents: 
47815diff
changeset | 140 | val attribute_global = attribute_generic o Context.Theory; | 
| 47815 | 141 | |
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 142 | fun attribute_cmd ctxt = attribute ctxt o check_src ctxt; | 
| 56032 
b034b9f0fa2a
clarified Args.check_src: retain name space information for error output;
 wenzelm parents: 
56029diff
changeset | 143 | fun attribute_cmd_global thy = attribute_global thy o check_src (Proof_Context.init_global thy); | 
| 18636 | 144 | |
| 5823 | 145 | |
| 17105 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 wenzelm parents: 
16934diff
changeset | 146 | (* attributed declarations *) | 
| 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 wenzelm parents: 
16934diff
changeset | 147 | |
| 45390 
e29521ef9059
tuned signature -- avoid spurious Thm.mixed_attribute;
 wenzelm parents: 
45375diff
changeset | 148 | fun map_specs f = map (apfst (apsnd f)); | 
| 30759 | 149 | |
| 45390 
e29521ef9059
tuned signature -- avoid spurious Thm.mixed_attribute;
 wenzelm parents: 
45375diff
changeset | 150 | fun map_facts f = map (apfst (apsnd f) o apsnd (map (apsnd f))); | 
| 30759 | 151 | fun map_facts_refs f g = map_facts f #> map (apsnd (map (apfst g))); | 
| 17105 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 wenzelm parents: 
16934diff
changeset | 152 | |
| 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 wenzelm parents: 
16934diff
changeset | 153 | |
| 38330 | 154 | (* fact expressions *) | 
| 155 | ||
| 47249 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 156 | fun global_notes kind facts thy = thy |> | 
| 47815 | 157 | Global_Theory.note_thmss kind (map_facts (map (attribute_global thy)) facts); | 
| 47249 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 158 | |
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 159 | fun local_notes kind facts ctxt = ctxt |> | 
| 47815 | 160 | Proof_Context.note_thmss kind (map_facts (map (attribute ctxt)) facts); | 
| 47249 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 161 | |
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 162 | fun generic_notes kind facts context = context |> | 
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 163 | Context.mapping_result (global_notes kind facts) (local_notes kind facts); | 
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 164 | |
| 38330 | 165 | fun eval_thms ctxt srcs = ctxt | 
| 42360 | 166 | |> Proof_Context.note_thmss "" | 
| 47815 | 167 | (map_facts_refs (map (attribute_cmd ctxt)) (Proof_Context.get_fact ctxt) | 
| 38330 | 168 | [((Binding.empty, []), srcs)]) | 
| 169 | |> fst |> maps snd; | |
| 170 | ||
| 171 | ||
| 30525 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 172 | (* attribute setup *) | 
| 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 173 | |
| 31306 
a74ee84288a0
eliminated old Attrib.add_attributes (and Attrib.syntax);
 wenzelm parents: 
31305diff
changeset | 174 | fun setup name scan = | 
| 
a74ee84288a0
eliminated old Attrib.add_attributes (and Attrib.syntax);
 wenzelm parents: 
31305diff
changeset | 175 | add_attribute name | 
| 56032 
b034b9f0fa2a
clarified Args.check_src: retain name space information for error output;
 wenzelm parents: 
56029diff
changeset | 176 | (fn src => fn (ctxt, th) => | 
| 
b034b9f0fa2a
clarified Args.check_src: retain name space information for error output;
 wenzelm parents: 
56029diff
changeset | 177 | let val (a, ctxt') = Args.syntax_generic scan src ctxt in a (ctxt', th) end); | 
| 5879 | 178 | |
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55740diff
changeset | 179 | fun attribute_setup name source cmt = | 
| 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55740diff
changeset | 180 | Context.theory_map (ML_Context.expression (#pos source) | 
| 30525 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 181 | "val (name, scan, comment): binding * attribute context_parser * string" | 
| 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 182 | "Context.map_theory (Attrib.setup name scan comment)" | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
36959diff
changeset | 183 |     (ML_Lex.read Position.none ("(" ^ ML_Syntax.make_binding name ^ ", ") @
 | 
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55740diff
changeset | 184 | ML_Lex.read_source source @ | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
36959diff
changeset | 185 |       ML_Lex.read Position.none (", " ^ ML_Syntax.print_string cmt ^ ")")));
 | 
| 30525 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 186 | |
| 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 187 | |
| 55140 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 188 | (* internal attribute *) | 
| 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 189 | |
| 56029 
8bedca4bd5a3
clarified Args.src: more abstract type, position refers to name only;
 wenzelm parents: 
56027diff
changeset | 190 | fun internal att = Args.src ("Pure.attribute", Position.none) [Token.mk_attribute att];
 | 
| 55140 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 191 | |
| 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 192 | val _ = Theory.setup | 
| 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 193 | (setup (Binding.name "attribute") (Scan.lift Args.internal_attribute >> Morphism.form) | 
| 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 194 | "internal attribute"); | 
| 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 195 | |
| 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 196 | |
| 31305 | 197 | (* add/del syntax *) | 
| 5823 | 198 | |
| 31305 | 199 | fun add_del add del = Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add); | 
| 5879 | 200 | |
| 201 | ||
| 25983 | 202 | |
| 203 | (** parsing attributed theorems **) | |
| 5879 | 204 | |
| 36950 | 205 | val thm_sel = Parse.$$$ "(" |-- Parse.list1
 | 
| 206 | (Parse.nat --| Parse.minus -- Parse.nat >> Facts.FromTo || | |
| 207 | Parse.nat --| Parse.minus >> Facts.From || | |
| 208 | Parse.nat >> Facts.Single) --| Parse.$$$ ")"; | |
| 27812 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 209 | |
| 18636 | 210 | local | 
| 211 | ||
| 21698 
43a842769765
thms etc.: proper treatment of internal_fact with selection;
 wenzelm parents: 
21658diff
changeset | 212 | val fact_name = Args.internal_fact >> K "<fact>" || Args.name; | 
| 
43a842769765
thms etc.: proper treatment of internal_fact with selection;
 wenzelm parents: 
21658diff
changeset | 213 | |
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 214 | fun gen_thm pick = Scan.depend (fn context => | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 215 | let | 
| 56140 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
56052diff
changeset | 216 | val get = Proof_Context.get_fact_generic context; | 
| 26361 | 217 | val get_fact = get o Facts.Fact; | 
| 27820 | 218 | fun get_named pos name = get (Facts.Named ((name, pos), NONE)); | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 219 | in | 
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 220 | Parse.$$$ "[" |-- Args.attribs (check_name_generic context) --| Parse.$$$ "]" >> (fn srcs => | 
| 24008 | 221 | let | 
| 47815 | 222 | val atts = map (attribute_generic context) srcs; | 
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46512diff
changeset | 223 | val (th', context') = fold (uncurry o Thm.apply_attribute) atts (Drule.dummy_thm, context); | 
| 55740 | 224 |       in (context', pick ("", Position.none) [th']) end)
 | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 225 | || | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 226 | (Scan.ahead Args.alt_name -- Args.named_fact get_fact | 
| 27820 | 227 |       >> (fn (s, fact) => ("", Facts.Fact s, fact)) ||
 | 
| 36950 | 228 | Scan.ahead (Parse.position fact_name) :|-- (fn (name, pos) => | 
| 27820 | 229 | Args.named_fact (get_named pos) -- Scan.option thm_sel | 
| 230 | >> (fn (fact, sel) => (name, Facts.Named ((name, pos), sel), fact)))) | |
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 231 | -- Args.opt_attribs (check_name_generic context) >> (fn ((name, thmref, fact), srcs) => | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 232 | let | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 233 | val ths = Facts.select thmref fact; | 
| 47815 | 234 | val atts = map (attribute_generic context) srcs; | 
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46512diff
changeset | 235 | val (ths', context') = | 
| 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46512diff
changeset | 236 | fold_map (curry (fold (uncurry o Thm.apply_attribute) atts)) ths context; | 
| 55740 | 237 | in (context', pick (name, Facts.pos_of_ref thmref) ths') end) | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 238 | end); | 
| 15456 
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
 berghofe parents: 
15117diff
changeset | 239 | |
| 18636 | 240 | in | 
| 241 | ||
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 242 | val thm = gen_thm Facts.the_single; | 
| 18998 | 243 | val multi_thm = gen_thm (K I); | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19473diff
changeset | 244 | val thms = Scan.repeat multi_thm >> flat; | 
| 18636 | 245 | |
| 246 | end; | |
| 247 | ||
| 5823 | 248 | |
| 5879 | 249 | |
| 45586 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 250 | (** partial evaluation -- observing rule/declaration/mixed attributes **) | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 251 | |
| 56140 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
56052diff
changeset | 252 | (*NB: result length may change due to rearrangement of symbolic expression*) | 
| 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
56052diff
changeset | 253 | |
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 254 | local | 
| 45526 | 255 | |
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 256 | fun apply_att src (context, th) = | 
| 45526 | 257 | let | 
| 55914 
c5b752d549e3
clarified init_assignable: make double-sure that initial values are reset;
 wenzelm parents: 
55828diff
changeset | 258 | val src1 = Args.init_assignable src; | 
| 47815 | 259 | val result = attribute_generic context src1 (context, th); | 
| 45526 | 260 | val src2 = Args.closure src1; | 
| 261 | in (src2, result) end; | |
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 262 | |
| 45527 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 263 | fun err msg src = | 
| 56029 
8bedca4bd5a3
clarified Args.src: more abstract type, position refers to name only;
 wenzelm parents: 
56027diff
changeset | 264 | let val (name, pos) = Args.name_of_src src | 
| 48992 | 265 | in error (msg ^ " " ^ quote name ^ Position.here pos) end; | 
| 45527 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 266 | |
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 267 | fun eval src ((th, dyn), (decls, context)) = | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 268 | (case (apply_att src (context, th), dyn) of | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 269 | ((_, (NONE, SOME th')), NONE) => ((th', NONE), (decls, context)) | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 270 | | ((_, (NONE, SOME _)), SOME _) => err "Mixed dynamic attribute followed by static rule" src | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 271 | | ((src', (SOME context', NONE)), NONE) => | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 272 | let | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 273 | val decls' = | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 274 | (case decls of | 
| 45526 | 275 | [] => [(th, [src'])] | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 276 | | (th2, srcs2) :: rest => | 
| 52683 
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
 wenzelm parents: 
52487diff
changeset | 277 | if Thm.eq_thm_strict (th, th2) | 
| 45526 | 278 | then ((th2, src' :: srcs2) :: rest) | 
| 279 | else (th, [src']) :: (th2, srcs2) :: rest); | |
| 45527 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 280 | in ((th, NONE), (decls', context')) end | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 281 | | ((src', (opt_context', opt_th')), _) => | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 282 | let | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 283 | val context' = the_default context opt_context'; | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 284 | val th' = the_default th opt_th'; | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 285 | val dyn' = | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 286 | (case dyn of | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 287 | NONE => SOME (th, [src']) | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 288 | | SOME (dyn_th, srcs) => SOME (dyn_th, src' :: srcs)); | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 289 | in ((th', dyn'), (decls, context')) end); | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 290 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 291 | in | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 292 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 293 | fun partial_evaluation ctxt facts = | 
| 55998 
f5f9fad3321c
keep current context visibility for PIDE markup and completion (in contrast to 8e3e004f1c31): Attrib.check_src of 9dc5ce83202c should intern/report attributes once, which happens for local_theory in the (visible) auxiliary context;
 wenzelm parents: 
55997diff
changeset | 294 | (facts, Context.Proof ctxt) |-> | 
| 45586 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 295 | fold_map (fn ((b, more_atts), fact) => fn context => | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 296 | let | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 297 | val (fact', (decls, context')) = | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 298 | (fact, ([], context)) |-> fold_map (fn (ths, atts) => fn res1 => | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 299 | (ths, res1) |-> fold_map (fn th => fn res2 => | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 300 | let | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 301 | val ((th', dyn'), res3) = fold eval (atts @ more_atts) ((th, NONE), res2); | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 302 | val th_atts' = | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 303 | (case dyn' of | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 304 | NONE => (th', []) | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 305 | | SOME (dyn_th', atts') => (dyn_th', rev atts')); | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 306 | in (th_atts', res3) end)) | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 307 | |>> flat; | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 308 | val decls' = rev (map (apsnd rev) decls); | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 309 | val facts' = | 
| 52683 
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
 wenzelm parents: 
52487diff
changeset | 310 | if eq_list (eq_fst Thm.eq_thm_strict) (decls', fact') then | 
| 45586 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 311 | [((b, []), map2 (fn (th, atts1) => fn (_, atts2) => (th, atts1 @ atts2)) decls' fact')] | 
| 45619 | 312 | else if null decls' then [((b, []), fact')] | 
| 45586 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 313 | else [(empty_binding, decls'), ((b, []), fact')]; | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 314 | in (facts', context') end) | 
| 46906 
3c1787d46935
suppress vacous notes elements, with subtle change of semantics: 'interpret' no longer pulls-in unnamed facts "by fact";
 wenzelm parents: 
46903diff
changeset | 315 | |> fst |> flat |> map (apsnd (map (apfst single))) | 
| 
3c1787d46935
suppress vacous notes elements, with subtle change of semantics: 'interpret' no longer pulls-in unnamed facts "by fact";
 wenzelm parents: 
46903diff
changeset | 316 | |> filter_out (fn (b, fact) => is_empty_binding b andalso forall (null o #2) fact); | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 317 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 318 | end; | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 319 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 320 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 321 | |
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 322 | (** configuration options **) | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 323 | |
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 324 | (* naming *) | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 325 | |
| 33522 | 326 | structure Configs = Theory_Data | 
| 24713 | 327 | ( | 
| 39163 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 wenzelm parents: 
39137diff
changeset | 328 | type T = Config.raw Symtab.table; | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 329 | val empty = Symtab.empty; | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 330 | val extend = I; | 
| 33522 | 331 | fun merge data = Symtab.merge (K true) data; | 
| 24713 | 332 | ); | 
| 5823 | 333 | |
| 52060 | 334 | fun print_options ctxt = | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 335 | let | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 336 | fun prt (name, config) = | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 337 | let val value = Config.get ctxt config in | 
| 50301 | 338 |         Pretty.block [Pretty.mark_str name, Pretty.str (": " ^ Config.print_type value ^ " ="),
 | 
| 339 | Pretty.brk 1, Pretty.str (Config.print_value value)] | |
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 340 | end; | 
| 56027 
25889f5c39a8
prefer Name_Space.pretty with its builtin markup;
 wenzelm parents: 
56026diff
changeset | 341 | val space = attribute_space ctxt; | 
| 56052 | 342 | val configs = | 
| 343 | Name_Space.markup_entries ctxt space | |
| 344 | (Symtab.dest (Configs.get (Proof_Context.theory_of ctxt))); | |
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 345 | in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end; | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 346 | |
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 347 | |
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 348 | (* concrete syntax *) | 
| 23988 
aa46577f4f44
added attribute "option" for setting configuration options;
 wenzelm parents: 
23937diff
changeset | 349 | |
| 24003 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 350 | local | 
| 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 351 | |
| 36950 | 352 | val equals = Parse.$$$ "="; | 
| 24003 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 353 | |
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 354 | fun scan_value (Config.Bool _) = | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 355 | equals -- Args.$$$ "false" >> K (Config.Bool false) || | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 356 | equals -- Args.$$$ "true" >> K (Config.Bool true) || | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 357 | Scan.succeed (Config.Bool true) | 
| 36950 | 358 | | scan_value (Config.Int _) = equals |-- Parse.int >> Config.Int | 
| 40291 | 359 | | scan_value (Config.Real _) = equals |-- Parse.real >> Config.Real | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 360 | | scan_value (Config.String _) = equals |-- Args.name >> Config.String; | 
| 24003 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 361 | |
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 362 | fun scan_config thy config = | 
| 36787 
f60e4dd6d76f
renamed Config.get_thy to Config.get_global etc. to indicate that this is not the real thing;
 wenzelm parents: 
36002diff
changeset | 363 | let val config_type = Config.get_global thy config | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 364 | in scan_value config_type >> (K o Thm.declaration_attribute o K o Config.put_generic config) end; | 
| 24003 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 365 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 366 | fun register binding config thy = | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 367 | let val name = Sign.full_name thy binding in | 
| 24126 
913e1fa904fb
turned simp_depth_limit into configuration option;
 wenzelm parents: 
24110diff
changeset | 368 | thy | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 369 | |> setup binding (Scan.lift (scan_config thy config) >> Morphism.form) "configuration option" | 
| 24126 
913e1fa904fb
turned simp_depth_limit into configuration option;
 wenzelm parents: 
24110diff
changeset | 370 | |> Configs.map (Symtab.update (name, config)) | 
| 
913e1fa904fb
turned simp_depth_limit into configuration option;
 wenzelm parents: 
24110diff
changeset | 371 | end; | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 372 | |
| 42808 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 373 | fun declare make coerce binding default = | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 374 | let | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 375 | val name = Binding.name_of binding; | 
| 52039 | 376 | val config_value = Config.declare name (make o default); | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 377 | val config = coerce config_value; | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 378 | in (config, register binding config_value) end; | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 379 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 380 | in | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 381 | |
| 52040 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 382 | fun register_config config = register (Binding.name (Config.name_of config)) config; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 383 | |
| 42808 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 384 | val config_bool = declare Config.Bool Config.bool; | 
| 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 385 | val config_int = declare Config.Int Config.int; | 
| 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 386 | val config_real = declare Config.Real Config.real; | 
| 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 387 | val config_string = declare Config.String Config.string; | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 388 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 389 | end; | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 390 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 391 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 392 | (* implicit setup *) | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 393 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 394 | local | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 395 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 396 | fun setup_config declare_config binding default = | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 397 | let | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 398 | val (config, setup) = declare_config binding default; | 
| 53171 | 399 | val _ = Theory.setup setup; | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 400 | in config end; | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 401 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 402 | in | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 403 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 404 | val setup_config_bool = setup_config config_bool; | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 405 | val setup_config_int = setup_config config_int; | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 406 | val setup_config_string = setup_config config_string; | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 407 | val setup_config_real = setup_config config_real; | 
| 24003 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 408 | |
| 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 409 | end; | 
| 23988 
aa46577f4f44
added attribute "option" for setting configuration options;
 wenzelm parents: 
23937diff
changeset | 410 | |
| 
aa46577f4f44
added attribute "option" for setting configuration options;
 wenzelm parents: 
23937diff
changeset | 411 | |
| 52040 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 412 | (* system options *) | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 413 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 414 | local | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 415 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 416 | fun declare_option coerce name = | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 417 | let | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 418 | val config = Config.declare_option name; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 419 | in (coerce config, register_config config) end; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 420 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 421 | fun setup_option coerce name = | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 422 | let | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 423 | val config = Config.declare_option name; | 
| 53171 | 424 | val _ = Theory.setup (register_config config); | 
| 52040 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 425 | in coerce config end; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 426 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 427 | in | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 428 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 429 | val option_bool = declare_option Config.bool; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 430 | val option_int = declare_option Config.int; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 431 | val option_real = declare_option Config.real; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 432 | val option_string = declare_option Config.string; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 433 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 434 | val setup_option_bool = setup_option Config.bool; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 435 | val setup_option_int = setup_option Config.int; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 436 | val setup_option_real = setup_option Config.real; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 437 | val setup_option_string = setup_option Config.string; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 438 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 439 | end; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 440 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 441 | |
| 18636 | 442 | (* theory setup *) | 
| 5823 | 443 | |
| 53171 | 444 | val _ = Theory.setup | 
| 52059 | 445 | (register_config quick_and_dirty_raw #> | 
| 446 | register_config Ast.trace_raw #> | |
| 43347 | 447 | register_config Ast.stats_raw #> | 
| 42289 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 448 | register_config Printer.show_brackets_raw #> | 
| 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 449 | register_config Printer.show_sorts_raw #> | 
| 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 450 | register_config Printer.show_types_raw #> | 
| 49657 
40e4feac2921
turn constraints into Isabelle_Markup.typing, depending on show_markup options;
 wenzelm parents: 
48992diff
changeset | 451 | register_config Printer.show_markup_raw #> | 
| 42289 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 452 | register_config Printer.show_structs_raw #> | 
| 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 453 | register_config Printer.show_question_marks_raw #> | 
| 46512 
4f9f61f9b535
simplified configuration options for syntax ambiguity;
 wenzelm parents: 
46506diff
changeset | 454 | register_config Syntax.ambiguity_warning_raw #> | 
| 46506 
c7faa011bfa7
simplified configuration options for syntax ambiguity;
 wenzelm parents: 
46126diff
changeset | 455 | register_config Syntax.ambiguity_limit_raw #> | 
| 42284 | 456 | register_config Syntax_Trans.eta_contract_raw #> | 
| 42669 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 wenzelm parents: 
42616diff
changeset | 457 | register_config Name_Space.names_long_raw #> | 
| 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 wenzelm parents: 
42616diff
changeset | 458 | register_config Name_Space.names_short_raw #> | 
| 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 wenzelm parents: 
42616diff
changeset | 459 | register_config Name_Space.names_unique_raw #> | 
| 41375 | 460 | register_config ML_Context.trace_raw #> | 
| 56265 
785569927666
discontinued Toplevel.debug in favour of system option "exception_trace";
 wenzelm parents: 
56140diff
changeset | 461 | register_config Runtime.exception_trace_raw #> | 
| 42360 | 462 | register_config Proof_Context.show_abbrevs_raw #> | 
| 39163 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 wenzelm parents: 
39137diff
changeset | 463 | register_config Goal_Display.goals_limit_raw #> | 
| 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 wenzelm parents: 
39137diff
changeset | 464 | register_config Goal_Display.show_main_goal_raw #> | 
| 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 wenzelm parents: 
39137diff
changeset | 465 | register_config Goal_Display.show_consts_raw #> | 
| 39166 
19efc2af3e6c
turned show_hyps and show_tags into proper configuration option;
 wenzelm parents: 
39163diff
changeset | 466 | register_config Display.show_hyps_raw #> | 
| 
19efc2af3e6c
turned show_hyps and show_tags into proper configuration option;
 wenzelm parents: 
39163diff
changeset | 467 | register_config Display.show_tags_raw #> | 
| 52709 
0e4bacf21e77
turned pattern unify flag into configuration option (global only);
 wenzelm parents: 
52683diff
changeset | 468 | register_config Pattern.unify_trace_failure_raw #> | 
| 39163 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 wenzelm parents: 
39137diff
changeset | 469 | register_config Unify.trace_bound_raw #> | 
| 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 wenzelm parents: 
39137diff
changeset | 470 | register_config Unify.search_bound_raw #> | 
| 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 wenzelm parents: 
39137diff
changeset | 471 | register_config Unify.trace_simp_raw #> | 
| 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 wenzelm parents: 
39137diff
changeset | 472 | register_config Unify.trace_types_raw #> | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41183diff
changeset | 473 | register_config Raw_Simplifier.simp_depth_limit_raw #> | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41183diff
changeset | 474 | register_config Raw_Simplifier.simp_trace_depth_limit_raw #> | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41183diff
changeset | 475 | register_config Raw_Simplifier.simp_debug_raw #> | 
| 53171 | 476 | register_config Raw_Simplifier.simp_trace_raw); | 
| 5823 | 477 | |
| 478 | end; |