| author | wenzelm | 
| Tue, 20 Dec 2016 10:06:18 +0100 | |
| changeset 64614 | 88211daacf93 | 
| parent 64556 | 851ae0e7b09c | 
| child 67147 | dea94b1aabc3 | 
| 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 | |
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59064diff
changeset | 9 | val print_attributes: bool -> Proof.context -> unit | 
| 63336 | 10 | val define_global: binding -> (Token.src -> attribute) -> string -> theory -> string * theory | 
| 11 | val define: binding -> (Token.src -> attribute) -> string -> local_theory -> string * local_theory | |
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 12 | 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 | 13 | val check_name: Proof.context -> xstring * Position.T -> string | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 14 | val check_src: Proof.context -> Token.src -> Token.src | 
| 61814 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 15 | val attribs: Token.src list context_parser | 
| 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 16 | val opt_attribs: Token.src list context_parser | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 17 | val pretty_attribs: Proof.context -> Token.src list -> Pretty.T list | 
| 63336 | 18 | val pretty_binding: Proof.context -> Attrib.binding -> string -> Pretty.T list | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 19 | val attribute: Proof.context -> Token.src -> attribute | 
| 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 20 | val attribute_global: theory -> Token.src -> attribute | 
| 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 21 | val attribute_cmd: Proof.context -> Token.src -> attribute | 
| 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 22 | val attribute_cmd_global: theory -> Token.src -> attribute | 
| 45390 
e29521ef9059
tuned signature -- avoid spurious Thm.mixed_attribute;
 wenzelm parents: 
45375diff
changeset | 23 |   val map_specs: ('a list -> 'att list) ->
 | 
| 30759 | 24 |     (('c * 'a list) * 'b) list -> (('c * 'att list) * 'b) list
 | 
| 45390 
e29521ef9059
tuned signature -- avoid spurious Thm.mixed_attribute;
 wenzelm parents: 
45375diff
changeset | 25 |   val map_facts: ('a list -> 'att list) ->
 | 
| 17105 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 wenzelm parents: 
16934diff
changeset | 26 |     (('c * 'a list) * ('d * 'a list) list) list ->
 | 
| 18905 | 27 |     (('c * 'att list) * ('d * 'att list) list) list
 | 
| 45390 
e29521ef9059
tuned signature -- avoid spurious Thm.mixed_attribute;
 wenzelm parents: 
45375diff
changeset | 28 |   val map_facts_refs: ('a list -> 'att list) -> ('b -> 'fact) ->
 | 
| 30759 | 29 |     (('c * 'a list) * ('b * 'a list) list) list ->
 | 
| 30 |     (('c * 'att list) * ('fact * 'att list) list) list
 | |
| 63267 | 31 | type thms = (thm list * Token.src list) list | 
| 63336 | 32 | val global_notes: string -> (Attrib.binding * thms) list -> | 
| 47249 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 33 | theory -> (string * thm list) list * theory | 
| 63336 | 34 | val local_notes: string -> (Attrib.binding * thms) list -> | 
| 47249 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 35 | Proof.context -> (string * thm list) list * Proof.context | 
| 63336 | 36 | val generic_notes: string -> (Attrib.binding * thms) list -> | 
| 47249 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 37 | Context.generic -> (string * thm list) list * Context.generic | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 38 | val eval_thms: Proof.context -> (Facts.ref * Token.src list) list -> thm list | 
| 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 39 | val attribute_syntax: attribute context_parser -> Token.src -> attribute | 
| 63336 | 40 | val setup: binding -> attribute context_parser -> string -> theory -> theory | 
| 41 | val local_setup: binding -> attribute context_parser -> string -> | |
| 57927 | 42 | local_theory -> local_theory | 
| 59064 | 43 | val attribute_setup: bstring * Position.T -> Input.source -> string -> | 
| 57941 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 44 | local_theory -> local_theory | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 45 | val internal: (morphism -> attribute) -> Token.src | 
| 63267 | 46 | val internal_declaration: declaration -> thms | 
| 30525 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 47 | val add_del: attribute -> attribute -> attribute context_parser | 
| 30513 | 48 | val thm: thm context_parser | 
| 49 | val thms: thm list context_parser | |
| 50 | val multi_thm: thm list context_parser | |
| 63267 | 51 | val transform_facts: morphism -> (Attrib.binding * thms) list -> (Attrib.binding * thms) list | 
| 63336 | 52 | val partial_evaluation: Proof.context -> | 
| 53 | (Attrib.binding * thms) list -> (Attrib.binding * thms) list | |
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59064diff
changeset | 54 | val print_options: bool -> Proof.context -> unit | 
| 63336 | 55 | val config_bool: binding -> (Context.generic -> bool) -> bool Config.T * (theory -> theory) | 
| 56 | val config_int: binding -> (Context.generic -> int) -> int Config.T * (theory -> theory) | |
| 57 | val config_real: binding -> (Context.generic -> real) -> real Config.T * (theory -> theory) | |
| 58 | val config_string: binding -> (Context.generic -> string) -> string Config.T * (theory -> theory) | |
| 59 | val setup_config_bool: binding -> (Context.generic -> bool) -> bool Config.T | |
| 60 | val setup_config_int: binding -> (Context.generic -> int) -> int Config.T | |
| 61 | val setup_config_real: binding -> (Context.generic -> real) -> real Config.T | |
| 62 | val setup_config_string: binding -> (Context.generic -> string) -> string Config.T | |
| 56438 | 63 | val option_bool: string * Position.T -> bool Config.T * (theory -> theory) | 
| 64 | val option_int: string * Position.T -> int Config.T * (theory -> theory) | |
| 65 | val option_real: string * Position.T -> real Config.T * (theory -> theory) | |
| 66 | val option_string: string * Position.T -> string Config.T * (theory -> theory) | |
| 67 | val setup_option_bool: string * Position.T -> bool Config.T | |
| 68 | val setup_option_int: string * Position.T -> int Config.T | |
| 69 | val setup_option_real: string * Position.T -> real Config.T | |
| 70 | val setup_option_string: string * Position.T -> string Config.T | |
| 63019 | 71 | val consumes: int -> Token.src | 
| 72 | val constraints: int -> Token.src | |
| 73 | val cases_open: Token.src | |
| 74 | val case_names: string list -> Token.src | |
| 75 | val case_conclusion: string * string list -> Token.src | |
| 5823 | 76 | end; | 
| 77 | ||
| 63336 | 78 | structure Attrib: sig type binding = Attrib.binding include ATTRIB end = | 
| 5823 | 79 | struct | 
| 80 | ||
| 63336 | 81 | type binding = Attrib.binding; | 
| 45584 
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
 wenzelm parents: 
45537diff
changeset | 82 | |
| 28084 
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 | |
| 57927 | 89 | structure Attributes = Generic_Data | 
| 22846 | 90 | ( | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 91 | type T = ((Token.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 | |
| 63090 
7aa9ac5165e4
common entity definitions within a global or local theory context;
 wenzelm parents: 
63068diff
changeset | 97 | val ops_attributes = {get_data = Attributes.get, put_data = Attributes.put};
 | 
| 57927 | 98 | |
| 63090 
7aa9ac5165e4
common entity definitions within a global or local theory context;
 wenzelm parents: 
63068diff
changeset | 99 | val get_attributes = Attributes.get o Context.Proof; | 
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 100 | |
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59064diff
changeset | 101 | fun print_attributes verbose ctxt = | 
| 22846 | 102 | let | 
| 57927 | 103 | val attribs = get_attributes ctxt; | 
| 50301 | 104 | fun prt_attr (name, (_, "")) = Pretty.mark_str name | 
| 42813 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 wenzelm parents: 
42808diff
changeset | 105 | | prt_attr (name, (_, comment)) = | 
| 50301 | 106 | Pretty.block | 
| 107 | (Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment); | |
| 22846 | 108 | in | 
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59064diff
changeset | 109 | [Pretty.big_list "attributes:" (map prt_attr (Name_Space.markup_table verbose ctxt attribs))] | 
| 56334 
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
 wenzelm parents: 
56303diff
changeset | 110 | |> Pretty.writeln_chunks | 
| 22846 | 111 | end; | 
| 7611 | 112 | |
| 57927 | 113 | val attribute_space = Name_Space.space_of_table o get_attributes; | 
| 114 | ||
| 115 | ||
| 116 | (* define *) | |
| 117 | ||
| 63090 
7aa9ac5165e4
common entity definitions within a global or local theory context;
 wenzelm parents: 
63068diff
changeset | 118 | fun define_global binding att comment = | 
| 
7aa9ac5165e4
common entity definitions within a global or local theory context;
 wenzelm parents: 
63068diff
changeset | 119 | Entity.define_global ops_attributes binding (att, comment); | 
| 56027 
25889f5c39a8
prefer Name_Space.pretty with its builtin markup;
 wenzelm parents: 
56026diff
changeset | 120 | |
| 57941 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 121 | fun define binding att comment = | 
| 63090 
7aa9ac5165e4
common entity definitions within a global or local theory context;
 wenzelm parents: 
63068diff
changeset | 122 | Entity.define ops_attributes binding (att, comment); | 
| 31306 
a74ee84288a0
eliminated old Attrib.add_attributes (and Attrib.syntax);
 wenzelm parents: 
31305diff
changeset | 123 | |
| 5823 | 124 | |
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 125 | (* check *) | 
| 15703 | 126 | |
| 57927 | 127 | fun check_name_generic context = #1 o Name_Space.check context (Attributes.get context); | 
| 56032 
b034b9f0fa2a
clarified Args.check_src: retain name space information for error output;
 wenzelm parents: 
56029diff
changeset | 128 | val check_name = check_name_generic o Context.Proof; | 
| 15703 | 129 | |
| 56033 | 130 | fun check_src ctxt src = | 
| 61814 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 131 | let | 
| 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 132 | val _ = | 
| 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 133 | if Token.checked_src src then () | 
| 62795 
063d2f23cdf6
removed redundant Position.set_range -- already done in Position.range;
 wenzelm parents: 
62498diff
changeset | 134 | else Context_Position.report ctxt (#1 (Token.range_of src)) Markup.language_attribute; | 
| 61814 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 135 | in #1 (Token.check_src ctxt get_attributes src) end; | 
| 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 136 | |
| 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 137 | val attribs = | 
| 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 138 | Args.context -- Scan.lift Parse.attribs | 
| 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 139 | >> (fn (ctxt, srcs) => map (check_src ctxt) srcs); | 
| 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 140 | |
| 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 141 | val opt_attribs = Scan.optional attribs []; | 
| 21031 | 142 | |
| 143 | ||
| 144 | (* pretty printing *) | |
| 145 | ||
| 146 | fun pretty_attribs _ [] = [] | |
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 147 | | pretty_attribs ctxt srcs = [Pretty.enum "," "[" "]" (map (Token.pretty_src ctxt) srcs)]; | 
| 21031 | 148 | |
| 60242 
3a8501876dba
tuned output -- avoid empty quites and extra breaks;
 wenzelm parents: 
59917diff
changeset | 149 | fun pretty_binding ctxt (b, atts) sep = | 
| 60243 | 150 | (case (Binding.is_empty b, null atts) of | 
| 151 | (true, true) => [] | |
| 152 | | (false, true) => [Pretty.block [Binding.pretty b, Pretty.str sep]] | |
| 153 | | (true, false) => [Pretty.block (pretty_attribs ctxt atts @ [Pretty.str sep])] | |
| 154 | | (false, false) => | |
| 155 | [Pretty.block | |
| 156 | (Binding.pretty b :: Pretty.brk 1 :: pretty_attribs ctxt atts @ [Pretty.str sep])]); | |
| 60242 
3a8501876dba
tuned output -- avoid empty quites and extra breaks;
 wenzelm parents: 
59917diff
changeset | 157 | |
| 15703 | 158 | |
| 18734 | 159 | (* get attributes *) | 
| 5823 | 160 | |
| 47816 
cd0dfb06b0c8
prefer Context_Position.report_generic, which observes is_visible flag and thus reduces number of echos;
 wenzelm parents: 
47815diff
changeset | 161 | fun attribute_generic context = | 
| 57927 | 162 | let val table = Attributes.get context | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 163 | in fn src => #1 (Name_Space.get table (#1 (Token.name_of_src src))) src end; | 
| 5823 | 164 | |
| 47816 
cd0dfb06b0c8
prefer Context_Position.report_generic, which observes is_visible flag and thus reduces number of echos;
 wenzelm parents: 
47815diff
changeset | 165 | 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 | 166 | val attribute_global = attribute_generic o Context.Theory; | 
| 47815 | 167 | |
| 55997 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 wenzelm parents: 
55914diff
changeset | 168 | 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 | 169 | fun attribute_cmd_global thy = attribute_global thy o check_src (Proof_Context.init_global thy); | 
| 18636 | 170 | |
| 5823 | 171 | |
| 17105 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 wenzelm parents: 
16934diff
changeset | 172 | (* attributed declarations *) | 
| 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 wenzelm parents: 
16934diff
changeset | 173 | |
| 45390 
e29521ef9059
tuned signature -- avoid spurious Thm.mixed_attribute;
 wenzelm parents: 
45375diff
changeset | 174 | fun map_specs f = map (apfst (apsnd f)); | 
| 30759 | 175 | |
| 45390 
e29521ef9059
tuned signature -- avoid spurious Thm.mixed_attribute;
 wenzelm parents: 
45375diff
changeset | 176 | fun map_facts f = map (apfst (apsnd f) o apsnd (map (apsnd f))); | 
| 30759 | 177 | 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 | 178 | |
| 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 wenzelm parents: 
16934diff
changeset | 179 | |
| 38330 | 180 | (* fact expressions *) | 
| 181 | ||
| 63267 | 182 | type thms = (thm list * Token.src list) list; | 
| 183 | ||
| 47249 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 184 | fun global_notes kind facts thy = thy |> | 
| 47815 | 185 | 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 | 186 | |
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 187 | fun local_notes kind facts ctxt = ctxt |> | 
| 47815 | 188 | 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 | 189 | |
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 190 | fun generic_notes kind facts context = context |> | 
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 191 | 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 | 192 | |
| 38330 | 193 | fun eval_thms ctxt srcs = ctxt | 
| 42360 | 194 | |> Proof_Context.note_thmss "" | 
| 63352 | 195 | (map_facts_refs | 
| 196 | (map (attribute_cmd ctxt)) (Proof_Context.get_fact ctxt) [(Binding.empty_atts, srcs)]) | |
| 38330 | 197 | |> fst |> maps snd; | 
| 198 | ||
| 199 | ||
| 30525 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 200 | (* attribute setup *) | 
| 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 201 | |
| 57927 | 202 | fun attribute_syntax scan src (context, th) = | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 203 | let val (a, context') = Token.syntax_generic scan src context in a (context', th) end; | 
| 57927 | 204 | |
| 205 | fun setup binding scan comment = define_global binding (attribute_syntax scan) comment #> snd; | |
| 206 | fun local_setup binding scan comment = define binding (attribute_syntax scan) comment #> snd; | |
| 5879 | 207 | |
| 58979 
162a4c2e97bc
more careful ML source positions, for improved PIDE markup;
 wenzelm parents: 
58978diff
changeset | 208 | fun attribute_setup name source comment = | 
| 
162a4c2e97bc
more careful ML source positions, for improved PIDE markup;
 wenzelm parents: 
58978diff
changeset | 209 | ML_Lex.read_source false source | 
| 59064 | 210 | |> ML_Context.expression (Input.range_of source) "parser" "Thm.attribute context_parser" | 
| 58979 
162a4c2e97bc
more careful ML source positions, for improved PIDE markup;
 wenzelm parents: 
58978diff
changeset | 211 |     ("Context.map_proof (Attrib.local_setup " ^ ML_Syntax.atomic (ML_Syntax.make_binding name) ^
 | 
| 
162a4c2e97bc
more careful ML source positions, for improved PIDE markup;
 wenzelm parents: 
58978diff
changeset | 212 | " parser " ^ ML_Syntax.print_string comment ^ ")") | 
| 57941 
57200bdc2aa7
localized command 'method_setup' and 'attribute_setup';
 wenzelm parents: 
57938diff
changeset | 213 | |> Context.proof_map; | 
| 30525 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 214 | |
| 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 215 | |
| 55140 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 216 | (* internal attribute *) | 
| 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 217 | |
| 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 218 | val _ = Theory.setup | 
| 64556 | 219 |   (setup (Binding.make ("attribute", \<^here>))
 | 
| 56436 | 220 | (Scan.lift Args.internal_attribute >> Morphism.form) | 
| 55140 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 221 | "internal attribute"); | 
| 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 222 | |
| 63019 | 223 | fun internal_name ctxt name = | 
| 224 | Token.make_src (name, Position.none) [] |> check_src ctxt |> hd; | |
| 225 | ||
| 226 | val internal_attribute_name = | |
| 227 | internal_name (Context.the_local_context ()) "attribute"; | |
| 228 | ||
| 229 | fun internal att = | |
| 230 | internal_attribute_name :: | |
| 231 |     [Token.make_string ("<attribute>", Position.none) |> Token.assign (SOME (Token.Attribute att))];
 | |
| 232 | ||
| 63266 | 233 | fun internal_declaration decl = | 
| 234 | [([Drule.dummy_thm], [internal (fn phi => Thm.declaration_attribute (K (decl phi)))])]; | |
| 235 | ||
| 55140 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 wenzelm parents: 
53171diff
changeset | 236 | |
| 31305 | 237 | (* add/del syntax *) | 
| 5823 | 238 | |
| 31305 | 239 | fun add_del add del = Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add); | 
| 5879 | 240 | |
| 241 | ||
| 25983 | 242 | |
| 243 | (** parsing attributed theorems **) | |
| 5879 | 244 | |
| 18636 | 245 | local | 
| 246 | ||
| 21698 
43a842769765
thms etc.: proper treatment of internal_fact with selection;
 wenzelm parents: 
21658diff
changeset | 247 | val fact_name = Args.internal_fact >> K "<fact>" || Args.name; | 
| 
43a842769765
thms etc.: proper treatment of internal_fact with selection;
 wenzelm parents: 
21658diff
changeset | 248 | |
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 249 | fun gen_thm pick = Scan.depend (fn context => | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 250 | let | 
| 56140 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
56052diff
changeset | 251 | val get = Proof_Context.get_fact_generic context; | 
| 26361 | 252 | val get_fact = get o Facts.Fact; | 
| 57942 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 wenzelm parents: 
57941diff
changeset | 253 | fun get_named is_sel pos name = | 
| 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 wenzelm parents: 
57941diff
changeset | 254 | let val (a, ths) = get (Facts.Named ((name, pos), NONE)) | 
| 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 wenzelm parents: 
57941diff
changeset | 255 | in (if is_sel then NONE else a, ths) end; | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 256 | in | 
| 61814 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 257 | Parse.$$$ "[" |-- Scan.pass context attribs --| Parse.$$$ "]" >> (fn srcs => | 
| 24008 | 258 | let | 
| 47815 | 259 | val atts = map (attribute_generic context) srcs; | 
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46512diff
changeset | 260 | val (th', context') = fold (uncurry o Thm.apply_attribute) atts (Drule.dummy_thm, context); | 
| 55740 | 261 |       in (context', pick ("", Position.none) [th']) end)
 | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 262 | || | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 263 | (Scan.ahead Args.alt_name -- Args.named_fact get_fact | 
| 27820 | 264 |       >> (fn (s, fact) => ("", Facts.Fact s, fact)) ||
 | 
| 58028 
e4250d370657
tuned signature -- define some elementary operations earlier;
 wenzelm parents: 
58025diff
changeset | 265 | Scan.ahead (Parse.position fact_name -- Scan.option Parse.thm_sel) :|-- | 
| 
e4250d370657
tuned signature -- define some elementary operations earlier;
 wenzelm parents: 
58025diff
changeset | 266 | (fn ((name, pos), sel) => | 
| 
e4250d370657
tuned signature -- define some elementary operations earlier;
 wenzelm parents: 
58025diff
changeset | 267 | Args.named_fact (get_named (is_some sel) pos) --| Scan.option Parse.thm_sel | 
| 
e4250d370657
tuned signature -- define some elementary operations earlier;
 wenzelm parents: 
58025diff
changeset | 268 | >> (fn fact => (name, Facts.Named ((name, pos), sel), fact)))) | 
| 61814 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 269 | -- Scan.pass context opt_attribs >> (fn ((name, thmref, fact), srcs) => | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 270 | let | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 271 | val ths = Facts.select thmref fact; | 
| 47815 | 272 | val atts = map (attribute_generic context) srcs; | 
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46512diff
changeset | 273 | val (ths', context') = | 
| 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46512diff
changeset | 274 | fold_map (curry (fold (uncurry o Thm.apply_attribute) atts)) ths context; | 
| 55740 | 275 | 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 | 276 | end); | 
| 15456 
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
 berghofe parents: 
15117diff
changeset | 277 | |
| 18636 | 278 | in | 
| 279 | ||
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 280 | val thm = gen_thm Facts.the_single; | 
| 18998 | 281 | val multi_thm = gen_thm (K I); | 
| 61476 | 282 | val thms = Scan.repeats multi_thm; | 
| 18636 | 283 | |
| 284 | end; | |
| 285 | ||
| 5823 | 286 | |
| 58028 
e4250d370657
tuned signature -- define some elementary operations earlier;
 wenzelm parents: 
58025diff
changeset | 287 | (* transform fact expressions *) | 
| 
e4250d370657
tuned signature -- define some elementary operations earlier;
 wenzelm parents: 
58025diff
changeset | 288 | |
| 
e4250d370657
tuned signature -- define some elementary operations earlier;
 wenzelm parents: 
58025diff
changeset | 289 | fun transform_facts phi = map (fn ((a, atts), bs) => | 
| 61814 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 290 | ((Morphism.binding phi a, (map o map) (Token.transform phi) atts), | 
| 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 291 | bs |> map (fn (ths, btts) => (Morphism.fact phi ths, (map o map) (Token.transform phi) btts)))); | 
| 58028 
e4250d370657
tuned signature -- define some elementary operations earlier;
 wenzelm parents: 
58025diff
changeset | 292 | |
| 
e4250d370657
tuned signature -- define some elementary operations earlier;
 wenzelm parents: 
58025diff
changeset | 293 | |
| 5879 | 294 | |
| 45586 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 295 | (** partial evaluation -- observing rule/declaration/mixed attributes **) | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 296 | |
| 56140 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
56052diff
changeset | 297 | (*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 | 298 | |
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 299 | local | 
| 45526 | 300 | |
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 301 | fun apply_att src (context, th) = | 
| 45526 | 302 | let | 
| 61814 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 303 | val src1 = map Token.init_assignable src; | 
| 47815 | 304 | val result = attribute_generic context src1 (context, th); | 
| 61814 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61527diff
changeset | 305 | val src2 = map Token.closure src1; | 
| 45526 | 306 | in (src2, result) end; | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 307 | |
| 45527 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 308 | fun err msg src = | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57942diff
changeset | 309 | let val (name, pos) = Token.name_of_src src | 
| 48992 | 310 | 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 | 311 | |
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 312 | fun eval src ((th, dyn), (decls, context)) = | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 313 | (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 | 314 | ((_, (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 | 315 | | ((_, (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 | 316 | | ((src', (SOME context', NONE)), NONE) => | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 317 | let | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 318 | val decls' = | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 319 | (case decls of | 
| 45526 | 320 | [] => [(th, [src'])] | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 321 | | (th2, srcs2) :: rest => | 
| 52683 
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
 wenzelm parents: 
52487diff
changeset | 322 | if Thm.eq_thm_strict (th, th2) | 
| 45526 | 323 | then ((th2, src' :: srcs2) :: rest) | 
| 324 | else (th, [src']) :: (th2, srcs2) :: rest); | |
| 45527 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 325 | in ((th, NONE), (decls', context')) end | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 326 | | ((src', (opt_context', opt_th')), _) => | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 327 | let | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 328 | val context' = the_default context opt_context'; | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 329 | val th' = the_default th opt_th'; | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 330 | val dyn' = | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 331 | (case dyn of | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 332 | NONE => SOME (th, [src']) | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 333 | | 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 | 334 | in ((th', dyn'), (decls, context')) end); | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 335 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 336 | in | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 337 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 338 | fun partial_evaluation ctxt facts = | 
| 57858 
39d9c7f175e0
refined context visibility again (amending f5f9fad3321c, 8e3e004f1c31): avoid spurious warning due to global config options;
 wenzelm parents: 
56438diff
changeset | 339 | (facts, Context.Proof (Context_Position.not_really ctxt)) |-> | 
| 45586 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 340 | 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 | 341 | let | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 342 | val (fact', (decls, context')) = | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 343 | (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 | 344 | (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 | 345 | let | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 346 | 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 | 347 | val th_atts' = | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 348 | (case dyn' of | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 349 | NONE => (th', []) | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 350 | | 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 | 351 | in (th_atts', res3) end)) | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 352 | |>> flat; | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 353 | val decls' = rev (map (apsnd rev) decls); | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 354 | val facts' = | 
| 52683 
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
 wenzelm parents: 
52487diff
changeset | 355 | 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 | 356 | [((b, []), map2 (fn (th, atts1) => fn (_, atts2) => (th, atts1 @ atts2)) decls' fact')] | 
| 45619 | 357 | else if null decls' then [((b, []), fact')] | 
| 63352 | 358 | else [(Binding.empty_atts, decls'), ((b, []), fact')]; | 
| 45586 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 359 | 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 | 360 | |> fst |> flat |> map (apsnd (map (apfst single))) | 
| 63352 | 361 | |> filter_out (fn (b, fact) => Binding.is_empty_atts b andalso forall (null o #2) fact); | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 362 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 363 | end; | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 364 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 365 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 366 | |
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 367 | (** configuration options **) | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 368 | |
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 369 | (* naming *) | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 370 | |
| 33522 | 371 | structure Configs = Theory_Data | 
| 24713 | 372 | ( | 
| 39163 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 wenzelm parents: 
39137diff
changeset | 373 | type T = Config.raw Symtab.table; | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 374 | val empty = Symtab.empty; | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 375 | val extend = I; | 
| 33522 | 376 | fun merge data = Symtab.merge (K true) data; | 
| 24713 | 377 | ); | 
| 5823 | 378 | |
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59064diff
changeset | 379 | fun print_options verbose ctxt = | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 380 | let | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 381 | fun prt (name, config) = | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 382 | let val value = Config.get ctxt config in | 
| 50301 | 383 |         Pretty.block [Pretty.mark_str name, Pretty.str (": " ^ Config.print_type value ^ " ="),
 | 
| 384 | Pretty.brk 1, Pretty.str (Config.print_value value)] | |
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 385 | end; | 
| 56027 
25889f5c39a8
prefer Name_Space.pretty with its builtin markup;
 wenzelm parents: 
56026diff
changeset | 386 | val space = attribute_space ctxt; | 
| 56052 | 387 | val configs = | 
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59064diff
changeset | 388 | Name_Space.markup_entries verbose ctxt space | 
| 56052 | 389 | (Symtab.dest (Configs.get (Proof_Context.theory_of ctxt))); | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 390 | in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end; | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 391 | |
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 392 | |
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 393 | (* concrete syntax *) | 
| 23988 
aa46577f4f44
added attribute "option" for setting configuration options;
 wenzelm parents: 
23937diff
changeset | 394 | |
| 24003 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 395 | local | 
| 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 396 | |
| 36950 | 397 | val equals = Parse.$$$ "="; | 
| 24003 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 398 | |
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 399 | fun scan_value (Config.Bool _) = | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 400 | equals -- Args.$$$ "false" >> K (Config.Bool false) || | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 401 | equals -- Args.$$$ "true" >> K (Config.Bool true) || | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 402 | Scan.succeed (Config.Bool true) | 
| 36950 | 403 | | scan_value (Config.Int _) = equals |-- Parse.int >> Config.Int | 
| 40291 | 404 | | scan_value (Config.Real _) = equals |-- Parse.real >> Config.Real | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 405 | | scan_value (Config.String _) = equals |-- Args.name >> Config.String; | 
| 24003 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 406 | |
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 407 | 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 | 408 | let val config_type = Config.get_global thy config | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 409 | 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 | 410 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 411 | fun register binding config thy = | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 412 | let val name = Sign.full_name thy binding in | 
| 24126 
913e1fa904fb
turned simp_depth_limit into configuration option;
 wenzelm parents: 
24110diff
changeset | 413 | thy | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 414 | |> 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 | 415 | |> Configs.map (Symtab.update (name, config)) | 
| 
913e1fa904fb
turned simp_depth_limit into configuration option;
 wenzelm parents: 
24110diff
changeset | 416 | end; | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 417 | |
| 42808 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 418 | fun declare make coerce binding default = | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 419 | let | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 420 | val name = Binding.name_of binding; | 
| 56438 | 421 | val pos = Binding.pos_of binding; | 
| 422 | val config_value = Config.declare (name, pos) (make o default); | |
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 423 | val config = coerce config_value; | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 424 | 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 | 425 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 426 | in | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 427 | |
| 56438 | 428 | fun register_config config = | 
| 429 | register (Binding.make (Config.name_of config, Config.pos_of config)) config; | |
| 52040 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 430 | |
| 42808 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 431 | val config_bool = declare Config.Bool Config.bool; | 
| 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 432 | val config_int = declare Config.Int Config.int; | 
| 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 433 | val config_real = declare Config.Real Config.real; | 
| 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 434 | 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 | 435 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 436 | end; | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 437 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 438 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 439 | (* implicit setup *) | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 440 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 441 | local | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 442 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 443 | 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 | 444 | let | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 445 | val (config, setup) = declare_config binding default; | 
| 53171 | 446 | val _ = Theory.setup setup; | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 447 | in config end; | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 448 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 449 | in | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 450 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 451 | 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 | 452 | 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 | 453 | 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 | 454 | val setup_config_real = setup_config config_real; | 
| 24003 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 455 | |
| 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 456 | end; | 
| 23988 
aa46577f4f44
added attribute "option" for setting configuration options;
 wenzelm parents: 
23937diff
changeset | 457 | |
| 
aa46577f4f44
added attribute "option" for setting configuration options;
 wenzelm parents: 
23937diff
changeset | 458 | |
| 52040 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 459 | (* system options *) | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 460 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 461 | local | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 462 | |
| 56438 | 463 | fun declare_option coerce (name, pos) = | 
| 52040 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 464 | let | 
| 56438 | 465 | val config = Config.declare_option (name, pos); | 
| 52040 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 466 | in (coerce config, register_config config) end; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 467 | |
| 56438 | 468 | fun setup_option coerce (name, pos) = | 
| 52040 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 469 | let | 
| 56438 | 470 | val config = Config.declare_option (name, pos); | 
| 53171 | 471 | val _ = Theory.setup (register_config config); | 
| 52040 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 472 | in coerce config end; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 473 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 474 | in | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 475 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 476 | val option_bool = declare_option Config.bool; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 477 | val option_int = declare_option Config.int; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 478 | val option_real = declare_option Config.real; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 479 | val option_string = declare_option Config.string; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 480 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 481 | 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 | 482 | 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 | 483 | 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 | 484 | 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 | 485 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 486 | end; | 
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 487 | |
| 
852939d19216
system options as context-sensitive configuration options within the attribute name space;
 wenzelm parents: 
52039diff
changeset | 488 | |
| 18636 | 489 | (* theory setup *) | 
| 5823 | 490 | |
| 53171 | 491 | val _ = Theory.setup | 
| 62898 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 492 |  (setup @{binding tagged} (Scan.lift (Args.name -- Args.name) >> Thm.tag) "tagged theorem" #>
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 493 |   setup @{binding untagged} (Scan.lift Args.name >> Thm.untag) "untagged theorem" #>
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 494 |   setup @{binding kind} (Scan.lift Args.name >> Thm.kind) "theorem kind" #>
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 495 |   setup @{binding THEN}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 496 | (Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- thm | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 497 | >> (fn (i, B) => Thm.rule_attribute [B] (fn _ => fn A => A RSN (i, B)))) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 498 | "resolution with rule" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 499 |   setup @{binding OF}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 500 | (thms >> (fn Bs => Thm.rule_attribute Bs (fn _ => fn A => A OF Bs))) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 501 | "rule resolved with facts" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 502 |   setup @{binding rename_abs}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 503 | (Scan.lift (Scan.repeat (Args.maybe Args.name)) >> (fn vs => | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 504 | Thm.rule_attribute [] (K (Drule.rename_bvars' vs)))) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 505 | "rename bound variables in abstractions" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 506 |   setup @{binding unfolded}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 507 | (thms >> (fn ths => | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 508 | Thm.rule_attribute ths (fn context => Local_Defs.unfold (Context.proof_of context) ths))) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 509 | "unfolded definitions" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 510 |   setup @{binding folded}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 511 | (thms >> (fn ths => | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 512 | Thm.rule_attribute ths (fn context => Local_Defs.fold (Context.proof_of context) ths))) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 513 | "folded definitions" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 514 |   setup @{binding consumes}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 515 | (Scan.lift (Scan.optional Parse.int 1) >> Rule_Cases.consumes) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 516 | "number of consumed facts" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 517 |   setup @{binding constraints}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 518 | (Scan.lift Parse.nat >> Rule_Cases.constraints) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 519 | "number of equality constraints" #> | 
| 63019 | 520 |   setup @{binding cases_open}
 | 
| 521 | (Scan.succeed Rule_Cases.cases_open) | |
| 522 | "rule with open parameters" #> | |
| 62898 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 523 |   setup @{binding case_names}
 | 
| 63019 | 524 | (Scan.lift (Scan.repeat (Args.name -- | 
| 62898 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 525 | Scan.optional (Parse.$$$ "[" |-- Scan.repeat1 (Args.maybe Args.name) --| Parse.$$$ "]") [])) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 526 | >> (fn cs => | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 527 | Rule_Cases.cases_hyp_names | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 528 | (map #1 cs) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 529 | (map (map (the_default Rule_Cases.case_hypsN) o #2) cs))) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 530 | "named rule cases" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 531 |   setup @{binding case_conclusion}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 532 | (Scan.lift (Args.name -- Scan.repeat Args.name) >> Rule_Cases.case_conclusion) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 533 | "named conclusion of rule cases" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 534 |   setup @{binding params}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 535 | (Scan.lift (Parse.and_list1 (Scan.repeat Args.name)) >> Rule_Cases.params) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 536 | "named rule parameters" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 537 |   setup @{binding rule_format}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 538 | (Scan.lift (Args.mode "no_asm") | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 539 | >> (fn true => Object_Logic.rule_format_no_asm | false => Object_Logic.rule_format)) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 540 | "result put into canonical rule format" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 541 |   setup @{binding elim_format}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 542 | (Scan.succeed (Thm.rule_attribute [] (K Tactic.make_elim))) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 543 | "destruct rule turned into elimination rule format" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 544 |   setup @{binding no_vars}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 545 | (Scan.succeed (Thm.rule_attribute [] (fn context => fn th => | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 546 | let | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 547 | val ctxt = Variable.set_body false (Context.proof_of context); | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 548 | val ((_, [th']), _) = Variable.import true [th] ctxt; | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 549 | in th' end))) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 550 | "imported schematic variables" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 551 |   setup @{binding atomize}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 552 | (Scan.succeed Object_Logic.declare_atomize) "declaration of atomize rule" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 553 |   setup @{binding rulify}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 554 | (Scan.succeed Object_Logic.declare_rulify) "declaration of rulify rule" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 555 |   setup @{binding rotated}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 556 | (Scan.lift (Scan.optional Parse.int 1 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 557 | >> (fn n => Thm.rule_attribute [] (fn _ => rotate_prems n)))) "rotated theorem premises" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 558 |   setup @{binding defn}
 | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 559 | (add_del Local_Defs.defn_add Local_Defs.defn_del) | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 560 | "declaration of definitional transformations" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 561 |   setup @{binding abs_def}
 | 
| 63068 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 wenzelm parents: 
63019diff
changeset | 562 | (Scan.succeed (Thm.rule_attribute [] (Local_Defs.abs_def_rule o Context.proof_of))) | 
| 62898 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 563 | "abstract over free variables of definitional theorem" #> | 
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 564 | |
| 
fdc290b68ecd
Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
 wenzelm parents: 
62878diff
changeset | 565 | register_config Goal.quick_and_dirty_raw #> | 
| 52059 | 566 | register_config Ast.trace_raw #> | 
| 43347 | 567 | register_config Ast.stats_raw #> | 
| 42289 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 568 | register_config Printer.show_brackets_raw #> | 
| 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 569 | register_config Printer.show_sorts_raw #> | 
| 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 570 | register_config Printer.show_types_raw #> | 
| 49657 
40e4feac2921
turn constraints into Isabelle_Markup.typing, depending on show_markup options;
 wenzelm parents: 
48992diff
changeset | 571 | register_config Printer.show_markup_raw #> | 
| 42289 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 572 | register_config Printer.show_structs_raw #> | 
| 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 573 | register_config Printer.show_question_marks_raw #> | 
| 46512 
4f9f61f9b535
simplified configuration options for syntax ambiguity;
 wenzelm parents: 
46506diff
changeset | 574 | register_config Syntax.ambiguity_warning_raw #> | 
| 46506 
c7faa011bfa7
simplified configuration options for syntax ambiguity;
 wenzelm parents: 
46126diff
changeset | 575 | register_config Syntax.ambiguity_limit_raw #> | 
| 42284 | 576 | register_config Syntax_Trans.eta_contract_raw #> | 
| 42669 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 wenzelm parents: 
42616diff
changeset | 577 | register_config Name_Space.names_long_raw #> | 
| 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 wenzelm parents: 
42616diff
changeset | 578 | register_config Name_Space.names_short_raw #> | 
| 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 wenzelm parents: 
42616diff
changeset | 579 | register_config Name_Space.names_unique_raw #> | 
| 62878 | 580 | register_config ML_Print_Depth.print_depth_raw #> | 
| 56303 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56281diff
changeset | 581 | register_config ML_Options.source_trace_raw #> | 
| 
4cc3f4db3447
clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
 wenzelm parents: 
56281diff
changeset | 582 | register_config ML_Options.exception_trace_raw #> | 
| 62498 | 583 | register_config ML_Options.exception_debugger_raw #> | 
| 60743 | 584 | register_config ML_Options.debugger_raw #> | 
| 42360 | 585 | 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 | 586 | 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 | 587 | register_config Goal_Display.show_main_goal_raw #> | 
| 61268 | 588 | register_config Thm.show_consts_raw #> | 
| 589 | register_config Thm.show_hyps_raw #> | |
| 590 | register_config Thm.show_tags_raw #> | |
| 52709 
0e4bacf21e77
turned pattern unify flag into configuration option (global only);
 wenzelm parents: 
52683diff
changeset | 591 | 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 | 592 | 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 | 593 | 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 | 594 | 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 | 595 | register_config Unify.trace_types_raw #> | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41183diff
changeset | 596 | register_config Raw_Simplifier.simp_depth_limit_raw #> | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41183diff
changeset | 597 | register_config Raw_Simplifier.simp_trace_depth_limit_raw #> | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41183diff
changeset | 598 | register_config Raw_Simplifier.simp_debug_raw #> | 
| 63068 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 wenzelm parents: 
63019diff
changeset | 599 | register_config Raw_Simplifier.simp_trace_raw #> | 
| 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 wenzelm parents: 
63019diff
changeset | 600 | register_config Local_Defs.unfold_abs_def_raw); | 
| 5823 | 601 | |
| 63019 | 602 | |
| 603 | (* internal source *) | |
| 604 | ||
| 605 | local | |
| 606 | ||
| 607 | val internal = internal_name (Context.the_local_context ()); | |
| 608 | ||
| 609 | val consumes_name = internal "consumes"; | |
| 610 | val constraints_name = internal "constraints"; | |
| 611 | val cases_open_name = internal "cases_open"; | |
| 612 | val case_names_name = internal "case_names"; | |
| 613 | val case_conclusion_name = internal "case_conclusion"; | |
| 614 | ||
| 615 | fun make_string s = Token.make_string (s, Position.none); | |
| 616 | ||
| 617 | in | |
| 618 | ||
| 619 | fun consumes i = consumes_name :: Token.make_int i; | |
| 620 | fun constraints i = constraints_name :: Token.make_int i; | |
| 621 | val cases_open = [cases_open_name]; | |
| 622 | fun case_names names = case_names_name :: map make_string names; | |
| 623 | fun case_conclusion (name, names) = case_conclusion_name :: map make_string (name :: names); | |
| 624 | ||
| 5823 | 625 | end; | 
| 63019 | 626 | |
| 627 | end; |