| author | wenzelm | 
| Thu, 11 Oct 2012 15:06:27 +0200 | |
| changeset 49817 | 85b37aece3b3 | 
| parent 49657 | 40e4feac2921 | 
| child 50301 | 56b4c9afd7be | 
| 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 | 
| 27729 | 13 | val print_attributes: theory -> unit | 
| 16458 | 14 | val intern: theory -> xstring -> string | 
| 15 | val intern_src: theory -> src -> src | |
| 21031 | 16 | val pretty_attribs: Proof.context -> src list -> Pretty.T list | 
| 26891 | 17 | val defined: theory -> string -> bool | 
| 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 | 
| 30575 | 38 | val attribute_setup: bstring * Position.T -> Symbol_Pos.text * Position.T -> string -> | 
| 39 | theory -> theory | |
| 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 | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 48 | val internal: (morphism -> attribute) -> src | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 49 | val print_configs: Proof.context -> unit | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 50 | val config_bool: Binding.binding -> | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 51 | (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 | 52 | val config_int: Binding.binding -> | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 53 | (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 | 54 | val config_real: Binding.binding -> | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 55 | (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 | 56 | val config_string: Binding.binding -> | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 57 | (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 | 58 | 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 | 59 | val setup_config_int: Binding.binding -> (Context.generic -> int) -> int Config.T | 
| 
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 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 61 | val setup_config_real: Binding.binding -> (Context.generic -> real) -> real Config.T | 
| 5823 | 62 | end; | 
| 63 | ||
| 64 | structure Attrib: ATTRIB = | |
| 65 | struct | |
| 66 | ||
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 67 | (* source and bindings *) | 
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 68 | |
| 15703 | 69 | type src = Args.src; | 
| 70 | ||
| 29581 | 71 | type binding = binding * src list; | 
| 45584 
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
 wenzelm parents: 
45537diff
changeset | 72 | |
| 28965 | 73 | val empty_binding: binding = (Binding.empty, []); | 
| 45584 
41a768a431a6
do not store vacuous theorem specifications -- relevant for frugal local theory content;
 wenzelm parents: 
45537diff
changeset | 74 | 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 | 75 | |
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 76 | |
| 27729 | 77 | |
| 18734 | 78 | (** named attributes **) | 
| 18636 | 79 | |
| 18734 | 80 | (* theory data *) | 
| 5823 | 81 | |
| 33522 | 82 | structure Attributes = Theory_Data | 
| 22846 | 83 | ( | 
| 33095 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 wenzelm parents: 
33092diff
changeset | 84 | type T = ((src -> attribute) * string) Name_Space.table; | 
| 33159 | 85 | val empty : T = Name_Space.empty_table "attribute"; | 
| 16458 | 86 | val extend = I; | 
| 33522 | 87 | fun merge data : T = Name_Space.merge_tables data; | 
| 22846 | 88 | ); | 
| 5823 | 89 | |
| 22846 | 90 | fun print_attributes thy = | 
| 91 | let | |
| 42360 | 92 | val ctxt = Proof_Context.init_global thy; | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 93 | val attribs = Attributes.get thy; | 
| 42813 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 wenzelm parents: 
42808diff
changeset | 94 | fun prt_attr (name, (_, "")) = Pretty.str name | 
| 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 wenzelm parents: 
42808diff
changeset | 95 | | prt_attr (name, (_, comment)) = | 
| 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 wenzelm parents: 
42808diff
changeset | 96 | Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; | 
| 22846 | 97 | in | 
| 42358 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 wenzelm parents: 
42289diff
changeset | 98 | [Pretty.big_list "attributes:" (map prt_attr (Name_Space.extern_table ctxt attribs))] | 
| 22846 | 99 | |> Pretty.chunks |> Pretty.writeln | 
| 100 | end; | |
| 7611 | 101 | |
| 33092 
c859019d3ac5
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 wenzelm parents: 
31794diff
changeset | 102 | 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 | 103 | |> 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 | 104 | |
| 5823 | 105 | |
| 21031 | 106 | (* name space *) | 
| 15703 | 107 | |
| 33095 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 wenzelm parents: 
33092diff
changeset | 108 | val intern = Name_Space.intern o #1 o Attributes.get; | 
| 15703 | 109 | val intern_src = Args.map_name o intern; | 
| 110 | ||
| 42360 | 111 | fun extern ctxt = Name_Space.extern ctxt (#1 (Attributes.get (Proof_Context.theory_of ctxt))); | 
| 21031 | 112 | |
| 113 | ||
| 114 | (* pretty printing *) | |
| 115 | ||
| 116 | fun pretty_attribs _ [] = [] | |
| 117 | | pretty_attribs ctxt srcs = | |
| 38329 | 118 | [Pretty.enum "," "[" "]" (map (Args.pretty_src ctxt o Args.map_name (extern ctxt)) srcs)]; | 
| 21031 | 119 | |
| 15703 | 120 | |
| 18734 | 121 | (* get attributes *) | 
| 5823 | 122 | |
| 26891 | 123 | val defined = Symtab.defined o #2 o Attributes.get; | 
| 124 | ||
| 47816 
cd0dfb06b0c8
prefer Context_Position.report_generic, which observes is_visible flag and thus reduces number of echos;
 wenzelm parents: 
47815diff
changeset | 125 | fun attribute_generic context = | 
| 5823 | 126 | let | 
| 47816 
cd0dfb06b0c8
prefer Context_Position.report_generic, which observes is_visible flag and thus reduces number of echos;
 wenzelm parents: 
47815diff
changeset | 127 | val thy = Context.theory_of context; | 
| 42380 | 128 | val (space, tab) = Attributes.get thy; | 
| 5879 | 129 | fun attr src = | 
| 16344 | 130 | let val ((name, _), pos) = Args.dest_src src in | 
| 42380 | 131 | (case Symtab.lookup tab name of | 
| 48992 | 132 |           NONE => error ("Unknown attribute: " ^ quote name ^ Position.here pos)
 | 
| 47816 
cd0dfb06b0c8
prefer Context_Position.report_generic, which observes is_visible flag and thus reduces number of echos;
 wenzelm parents: 
47815diff
changeset | 133 | | SOME (att, _) => | 
| 
cd0dfb06b0c8
prefer Context_Position.report_generic, which observes is_visible flag and thus reduces number of echos;
 wenzelm parents: 
47815diff
changeset | 134 | (Context_Position.report_generic context pos (Name_Space.markup space name); att src)) | 
| 5823 | 135 | end; | 
| 136 | in attr end; | |
| 137 | ||
| 47816 
cd0dfb06b0c8
prefer Context_Position.report_generic, which observes is_visible flag and thus reduces number of echos;
 wenzelm parents: 
47815diff
changeset | 138 | 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 | 139 | val attribute_global = attribute_generic o Context.Theory; | 
| 47815 | 140 | |
| 47816 
cd0dfb06b0c8
prefer Context_Position.report_generic, which observes is_visible flag and thus reduces number of echos;
 wenzelm parents: 
47815diff
changeset | 141 | fun attribute_cmd ctxt = attribute ctxt o intern_src (Proof_Context.theory_of ctxt); | 
| 47815 | 142 | fun attribute_cmd_global thy = attribute_global thy o intern_src thy; | 
| 18636 | 143 | |
| 5823 | 144 | |
| 17105 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 wenzelm parents: 
16934diff
changeset | 145 | (* attributed declarations *) | 
| 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 wenzelm parents: 
16934diff
changeset | 146 | |
| 45390 
e29521ef9059
tuned signature -- avoid spurious Thm.mixed_attribute;
 wenzelm parents: 
45375diff
changeset | 147 | fun map_specs f = map (apfst (apsnd f)); | 
| 30759 | 148 | |
| 45390 
e29521ef9059
tuned signature -- avoid spurious Thm.mixed_attribute;
 wenzelm parents: 
45375diff
changeset | 149 | fun map_facts f = map (apfst (apsnd f) o apsnd (map (apsnd f))); | 
| 30759 | 150 | 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 | 151 | |
| 
1b07a176a880
added map_specs/facts operators (from locale.ML);
 wenzelm parents: 
16934diff
changeset | 152 | |
| 38330 | 153 | (* fact expressions *) | 
| 154 | ||
| 47249 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 155 | fun global_notes kind facts thy = thy |> | 
| 47815 | 156 | 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 | 157 | |
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 158 | fun local_notes kind facts ctxt = ctxt |> | 
| 47815 | 159 | 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 | 160 | |
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 161 | fun generic_notes kind facts context = context |> | 
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47005diff
changeset | 162 | 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 | 163 | |
| 38330 | 164 | fun eval_thms ctxt srcs = ctxt | 
| 42360 | 165 | |> Proof_Context.note_thmss "" | 
| 47815 | 166 | (map_facts_refs (map (attribute_cmd ctxt)) (Proof_Context.get_fact ctxt) | 
| 38330 | 167 | [((Binding.empty, []), srcs)]) | 
| 168 | |> fst |> maps snd; | |
| 169 | ||
| 170 | ||
| 30525 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 171 | (* attribute setup *) | 
| 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 172 | |
| 31306 
a74ee84288a0
eliminated old Attrib.add_attributes (and Attrib.syntax);
 wenzelm parents: 
31305diff
changeset | 173 | fun syntax scan = Args.syntax "attribute" scan; | 
| 30525 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 174 | |
| 31306 
a74ee84288a0
eliminated old Attrib.add_attributes (and Attrib.syntax);
 wenzelm parents: 
31305diff
changeset | 175 | fun setup name scan = | 
| 
a74ee84288a0
eliminated old Attrib.add_attributes (and Attrib.syntax);
 wenzelm parents: 
31305diff
changeset | 176 | add_attribute name | 
| 
a74ee84288a0
eliminated old Attrib.add_attributes (and Attrib.syntax);
 wenzelm parents: 
31305diff
changeset | 177 | (fn src => fn (ctxt, th) => let val (a, ctxt') = syntax scan src ctxt in a (ctxt', th) end); | 
| 5879 | 178 | |
| 30525 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 179 | fun attribute_setup name (txt, pos) cmt = | 
| 
8a5a0aa30e1c
added setup and attribute_setup -- expect plain parser instead of syntax function;
 wenzelm parents: 
30513diff
changeset | 180 | Context.theory_map (ML_Context.expression pos | 
| 
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 ^ ", ") @
 | 
| 
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 | 184 | ML_Lex.read pos txt @ | 
| 
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 | |
| 31305 | 188 | (* add/del syntax *) | 
| 5823 | 189 | |
| 31305 | 190 | fun add_del add del = Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add); | 
| 5879 | 191 | |
| 192 | ||
| 25983 | 193 | |
| 194 | (** parsing attributed theorems **) | |
| 5879 | 195 | |
| 36950 | 196 | val thm_sel = Parse.$$$ "(" |-- Parse.list1
 | 
| 197 | (Parse.nat --| Parse.minus -- Parse.nat >> Facts.FromTo || | |
| 198 | Parse.nat --| Parse.minus >> Facts.From || | |
| 199 | Parse.nat >> Facts.Single) --| Parse.$$$ ")"; | |
| 27812 
af8edf3ab68c
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27751diff
changeset | 200 | |
| 18636 | 201 | local | 
| 202 | ||
| 21698 
43a842769765
thms etc.: proper treatment of internal_fact with selection;
 wenzelm parents: 
21658diff
changeset | 203 | val fact_name = Args.internal_fact >> K "<fact>" || Args.name; | 
| 
43a842769765
thms etc.: proper treatment of internal_fact with selection;
 wenzelm parents: 
21658diff
changeset | 204 | |
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 205 | fun gen_thm pick = Scan.depend (fn context => | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 206 | let | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 207 | val thy = Context.theory_of context; | 
| 42360 | 208 | val get = Context.cases (Global_Theory.get_fact context) Proof_Context.get_fact context; | 
| 26361 | 209 | val get_fact = get o Facts.Fact; | 
| 27820 | 210 | 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 | 211 | in | 
| 36950 | 212 | Parse.$$$ "[" |-- Args.attribs (intern thy) --| Parse.$$$ "]" >> (fn srcs => | 
| 24008 | 213 | let | 
| 47815 | 214 | val atts = map (attribute_generic context) srcs; | 
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46512diff
changeset | 215 | val (th', context') = fold (uncurry o Thm.apply_attribute) atts (Drule.dummy_thm, context); | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 216 | in (context', pick "" [th']) end) | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 217 | || | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 218 | (Scan.ahead Args.alt_name -- Args.named_fact get_fact | 
| 27820 | 219 |       >> (fn (s, fact) => ("", Facts.Fact s, fact)) ||
 | 
| 36950 | 220 | Scan.ahead (Parse.position fact_name) :|-- (fn (name, pos) => | 
| 27820 | 221 | Args.named_fact (get_named pos) -- Scan.option thm_sel | 
| 222 | >> (fn (fact, sel) => (name, Facts.Named ((name, pos), sel), fact)))) | |
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 223 | -- Args.opt_attribs (intern thy) >> (fn ((name, thmref, fact), srcs) => | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 224 | let | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 225 | val ths = Facts.select thmref fact; | 
| 47815 | 226 | val atts = map (attribute_generic context) srcs; | 
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46512diff
changeset | 227 | val (ths', context') = | 
| 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46512diff
changeset | 228 | fold_map (curry (fold (uncurry o Thm.apply_attribute) atts)) ths context; | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 229 | in (context', pick name ths') end) | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 230 | end); | 
| 15456 
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
 berghofe parents: 
15117diff
changeset | 231 | |
| 18636 | 232 | in | 
| 233 | ||
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25983diff
changeset | 234 | val thm = gen_thm Facts.the_single; | 
| 18998 | 235 | val multi_thm = gen_thm (K I); | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19473diff
changeset | 236 | val thms = Scan.repeat multi_thm >> flat; | 
| 18636 | 237 | |
| 238 | end; | |
| 239 | ||
| 5823 | 240 | |
| 5879 | 241 | |
| 45586 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 242 | (** partial evaluation -- observing rule/declaration/mixed attributes **) | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 243 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 244 | local | 
| 45526 | 245 | |
| 45586 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 246 | val strict_eq_thm = op = o pairself Thm.rep_thm; | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 247 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 248 | fun apply_att src (context, th) = | 
| 45526 | 249 | let | 
| 250 | val src1 = Args.assignable src; | |
| 47815 | 251 | val result = attribute_generic context src1 (context, th); | 
| 45526 | 252 | val src2 = Args.closure src1; | 
| 253 | in (src2, result) end; | |
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 254 | |
| 45527 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 255 | fun err msg src = | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 256 | let val ((name, _), pos) = Args.dest_src src | 
| 48992 | 257 | 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 | 258 | |
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 259 | fun eval src ((th, dyn), (decls, context)) = | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 260 | (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 | 261 | ((_, (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 | 262 | | ((_, (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 | 263 | | ((src', (SOME context', NONE)), NONE) => | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 264 | let | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 265 | val decls' = | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 266 | (case decls of | 
| 45526 | 267 | [] => [(th, [src'])] | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 268 | | (th2, srcs2) :: rest => | 
| 45586 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 269 | if strict_eq_thm (th, th2) | 
| 45526 | 270 | then ((th2, src' :: srcs2) :: rest) | 
| 271 | else (th, [src']) :: (th2, srcs2) :: rest); | |
| 45527 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 272 | in ((th, NONE), (decls', context')) end | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 273 | | ((src', (opt_context', opt_th')), _) => | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 274 | let | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 275 | val context' = the_default context opt_context'; | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 276 | val th' = the_default th opt_th'; | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 277 | val dyn' = | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 278 | (case dyn of | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 279 | NONE => SOME (th, [src']) | 
| 
d2b3d16b673a
retain mixed attributes as dynamic theorem expression, but disallow subsequent static rules;
 wenzelm parents: 
45526diff
changeset | 280 | | 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 | 281 | in ((th', dyn'), (decls, context')) end); | 
| 45493 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 282 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 283 | in | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 284 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 285 | fun partial_evaluation ctxt facts = | 
| 45586 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 286 | (facts, Context.Proof (Context_Position.set_visible false ctxt)) |-> | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 287 | 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 | 288 | let | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 289 | val (fact', (decls, context')) = | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 290 | (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 | 291 | (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 | 292 | let | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 293 | 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 | 294 | val th_atts' = | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 295 | (case dyn' of | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 296 | NONE => (th', []) | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 297 | | 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 | 298 | in (th_atts', res3) end)) | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 299 | |>> flat; | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 300 | val decls' = rev (map (apsnd rev) decls); | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 301 | val facts' = | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 302 | if eq_list (eq_fst strict_eq_thm) (decls', fact') then | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 303 | [((b, []), map2 (fn (th, atts1) => fn (_, atts2) => (th, atts1 @ atts2)) decls' fact')] | 
| 45619 | 304 | 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 | 305 | else [(empty_binding, decls'), ((b, []), fact')]; | 
| 
c94f149cdf5d
refined partial evaluation of attributes: avoid duplication of facts for plain declarations;
 wenzelm parents: 
45584diff
changeset | 306 | 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 | 307 | |> 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 | 308 | |> 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 | 309 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 310 | end; | 
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 311 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 312 | |
| 
12453fd8dcff
some support for partial evaluation of attributed facts;
 wenzelm parents: 
45491diff
changeset | 313 | |
| 25983 | 314 | (** basic attributes **) | 
| 315 | ||
| 316 | (* internal *) | |
| 317 | ||
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36950diff
changeset | 318 | fun internal att = Args.src (("Pure.attribute", [Token.mk_attribute att]), Position.none);
 | 
| 25983 | 319 | |
| 320 | ||
| 321 | (* rule composition *) | |
| 322 | ||
| 323 | val THEN_att = | |
| 36950 | 324 | Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- thm | 
| 31305 | 325 | >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => A RSN (i, B))); | 
| 25983 | 326 | |
| 327 | val OF_att = | |
| 31305 | 328 | thms >> (fn Bs => Thm.rule_attribute (fn _ => fn A => Bs MRS A)); | 
| 25983 | 329 | |
| 330 | ||
| 331 | (* rename_abs *) | |
| 332 | ||
| 45375 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
44069diff
changeset | 333 | val rename_abs = | 
| 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
44069diff
changeset | 334 | Scan.repeat (Args.maybe Args.name) | 
| 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
44069diff
changeset | 335 | >> (fn args => Thm.rule_attribute (K (Drule.rename_bvars' args))); | 
| 25983 | 336 | |
| 337 | ||
| 338 | (* unfold / fold definitions *) | |
| 339 | ||
| 340 | fun unfolded_syntax rule = | |
| 31305 | 341 | thms >> (fn ths => Thm.rule_attribute (fn context => rule (Context.proof_of context) ths)); | 
| 25983 | 342 | |
| 35624 | 343 | val unfolded = unfolded_syntax Local_Defs.unfold; | 
| 344 | val folded = unfolded_syntax Local_Defs.fold; | |
| 25983 | 345 | |
| 346 | ||
| 347 | (* rule format *) | |
| 348 | ||
| 31305 | 349 | val rule_format = Args.mode "no_asm" | 
| 35625 | 350 | >> (fn true => Object_Logic.rule_format_no_asm | false => Object_Logic.rule_format); | 
| 25983 | 351 | |
| 31305 | 352 | val elim_format = Thm.rule_attribute (K Tactic.make_elim); | 
| 25983 | 353 | |
| 354 | ||
| 44046 
a43ca8ed6564
extended user-level attribute case_names with names for case hypotheses
 nipkow parents: 
43347diff
changeset | 355 | (* case names *) | 
| 
a43ca8ed6564
extended user-level attribute case_names with names for case hypotheses
 nipkow parents: 
43347diff
changeset | 356 | |
| 45375 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
44069diff
changeset | 357 | val case_names = | 
| 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
44069diff
changeset | 358 | Scan.repeat1 (Args.name -- | 
| 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
44069diff
changeset | 359 | Scan.optional (Parse.$$$ "[" |-- Scan.repeat1 (Args.maybe Args.name) --| Parse.$$$ "]") []) >> | 
| 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
44069diff
changeset | 360 | (fn cs => | 
| 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
44069diff
changeset | 361 | Rule_Cases.cases_hyp_names (map fst cs) | 
| 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
44069diff
changeset | 362 | (map (map (the_default Rule_Cases.case_hypsN) o snd) cs)); | 
| 44053 | 363 | |
| 44046 
a43ca8ed6564
extended user-level attribute case_names with names for case hypotheses
 nipkow parents: 
43347diff
changeset | 364 | |
| 25983 | 365 | (* misc rules *) | 
| 366 | ||
| 31305 | 367 | val no_vars = Thm.rule_attribute (fn context => fn th => | 
| 26715 
00ff79ab6e6b
no_vars: reset body mode, i.e. invent global frees (which are acceptable to Variable.auto_fixes);
 wenzelm parents: 
26685diff
changeset | 368 | let | 
| 
00ff79ab6e6b
no_vars: reset body mode, i.e. invent global frees (which are acceptable to Variable.auto_fixes);
 wenzelm parents: 
26685diff
changeset | 369 | val ctxt = Variable.set_body false (Context.proof_of context); | 
| 31794 
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
 wenzelm parents: 
31365diff
changeset | 370 | val ((_, [th']), _) = Variable.import true [th] ctxt; | 
| 31305 | 371 | in th' end); | 
| 25983 | 372 | |
| 373 | val eta_long = | |
| 31305 | 374 | Thm.rule_attribute (K (Conv.fconv_rule Drule.eta_long_conversion)); | 
| 25983 | 375 | |
| 36950 | 376 | val rotated = Scan.optional Parse.int 1 >> (fn n => Thm.rule_attribute (K (rotate_prems n))); | 
| 29690 | 377 | |
| 25983 | 378 | |
| 379 | (* theory setup *) | |
| 5823 | 380 | |
| 26463 | 381 | val _ = Context.>> (Context.map_theory | 
| 31305 | 382 | (setup (Binding.name "attribute") (Scan.lift Args.internal_attribute >> Morphism.form) | 
| 383 | "internal attribute" #> | |
| 384 | setup (Binding.name "tagged") (Scan.lift (Args.name -- Args.name) >> Thm.tag) "tagged theorem" #> | |
| 385 | setup (Binding.name "untagged") (Scan.lift Args.name >> Thm.untag) "untagged theorem" #> | |
| 386 | setup (Binding.name "kind") (Scan.lift Args.name >> Thm.kind) "theorem kind" #> | |
| 387 | setup (Binding.name "THEN") THEN_att "resolution with rule" #> | |
| 388 | setup (Binding.name "OF") OF_att "rule applied to facts" #> | |
| 389 | setup (Binding.name "rename_abs") (Scan.lift rename_abs) | |
| 390 | "rename bound variables in abstractions" #> | |
| 391 | setup (Binding.name "unfolded") unfolded "unfolded definitions" #> | |
| 392 | setup (Binding.name "folded") folded "folded definitions" #> | |
| 36950 | 393 | setup (Binding.name "consumes") (Scan.lift (Scan.optional Parse.nat 1) >> Rule_Cases.consumes) | 
| 31305 | 394 | "number of consumed facts" #> | 
| 36950 | 395 | setup (Binding.name "constraints") (Scan.lift Parse.nat >> Rule_Cases.constraints) | 
| 34986 
7f7939c9370f
Added "constraints" tag / attribute for specifying the number of equality
 berghofe parents: 
33666diff
changeset | 396 | "number of equality constraints" #> | 
| 45375 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
44069diff
changeset | 397 | setup (Binding.name "case_names") (Scan.lift case_names) "named rule cases" #> | 
| 31305 | 398 | setup (Binding.name "case_conclusion") | 
| 33368 | 399 | (Scan.lift (Args.name -- Scan.repeat Args.name) >> Rule_Cases.case_conclusion) | 
| 31305 | 400 | "named conclusion of rule cases" #> | 
| 401 | setup (Binding.name "params") | |
| 36950 | 402 | (Scan.lift (Parse.and_list1 (Scan.repeat Args.name)) >> Rule_Cases.params) | 
| 31305 | 403 | "named rule parameters" #> | 
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34986diff
changeset | 404 | setup (Binding.name "standard") (Scan.succeed (Thm.rule_attribute (K Drule.export_without_context))) | 
| 31305 | 405 | "result put into standard form (legacy)" #> | 
| 406 | setup (Binding.name "rule_format") rule_format "result put into canonical rule format" #> | |
| 407 | setup (Binding.name "elim_format") (Scan.succeed elim_format) | |
| 408 | "destruct rule turned into elimination rule format" #> | |
| 409 | setup (Binding.name "no_vars") (Scan.succeed no_vars) "frozen schematic vars" #> | |
| 410 | setup (Binding.name "eta_long") (Scan.succeed eta_long) | |
| 411 | "put theorem into eta long beta normal form" #> | |
| 35625 | 412 | setup (Binding.name "atomize") (Scan.succeed Object_Logic.declare_atomize) | 
| 31305 | 413 | "declaration of atomize rule" #> | 
| 35625 | 414 | setup (Binding.name "rulify") (Scan.succeed Object_Logic.declare_rulify) | 
| 31305 | 415 | "declaration of rulify rule" #> | 
| 416 | setup (Binding.name "rotated") (Scan.lift rotated) "rotated theorem premises" #> | |
| 35624 | 417 | setup (Binding.name "defn") (add_del Local_Defs.defn_add Local_Defs.defn_del) | 
| 31305 | 418 | "declaration of definitional transformations" #> | 
| 46903 
3d44892ac0d6
improved attribute "abs_def" to handle object-equality as well;
 wenzelm parents: 
46775diff
changeset | 419 | setup (Binding.name "abs_def") | 
| 
3d44892ac0d6
improved attribute "abs_def" to handle object-equality as well;
 wenzelm parents: 
46775diff
changeset | 420 | (Scan.succeed (Thm.rule_attribute (fn context => | 
| 
3d44892ac0d6
improved attribute "abs_def" to handle object-equality as well;
 wenzelm parents: 
46775diff
changeset | 421 | Local_Defs.meta_rewrite_rule (Context.proof_of context) #> Drule.abs_def))) | 
| 
3d44892ac0d6
improved attribute "abs_def" to handle object-equality as well;
 wenzelm parents: 
46775diff
changeset | 422 | "abstract over free variables of definitionial theorem")); | 
| 8633 | 423 | |
| 5823 | 424 | |
| 425 | ||
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 426 | (** configuration options **) | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 427 | |
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 428 | (* naming *) | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 429 | |
| 33522 | 430 | structure Configs = Theory_Data | 
| 24713 | 431 | ( | 
| 39163 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 wenzelm parents: 
39137diff
changeset | 432 | type T = Config.raw Symtab.table; | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 433 | val empty = Symtab.empty; | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 434 | val extend = I; | 
| 33522 | 435 | fun merge data = Symtab.merge (K true) data; | 
| 24713 | 436 | ); | 
| 5823 | 437 | |
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 438 | fun print_configs ctxt = | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 439 | let | 
| 42360 | 440 | val thy = Proof_Context.theory_of ctxt; | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 441 | fun prt (name, config) = | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 442 | let val value = Config.get ctxt config in | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 443 | Pretty.block [Pretty.str (name ^ ": " ^ Config.print_type value ^ " ="), Pretty.brk 1, | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 444 | Pretty.str (Config.print_value value)] | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 445 | end; | 
| 42358 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 wenzelm parents: 
42289diff
changeset | 446 | val configs = Name_Space.extern_table ctxt (#1 (Attributes.get thy), Configs.get thy); | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 447 | in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end; | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 448 | |
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 449 | |
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 450 | (* concrete syntax *) | 
| 23988 
aa46577f4f44
added attribute "option" for setting configuration options;
 wenzelm parents: 
23937diff
changeset | 451 | |
| 24003 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 452 | local | 
| 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 453 | |
| 36950 | 454 | val equals = Parse.$$$ "="; | 
| 24003 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 455 | |
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 456 | fun scan_value (Config.Bool _) = | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 457 | equals -- Args.$$$ "false" >> K (Config.Bool false) || | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 458 | equals -- Args.$$$ "true" >> K (Config.Bool true) || | 
| 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 459 | Scan.succeed (Config.Bool true) | 
| 36950 | 460 | | scan_value (Config.Int _) = equals |-- Parse.int >> Config.Int | 
| 40291 | 461 | | scan_value (Config.Real _) = equals |-- Parse.real >> Config.Real | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 462 | | scan_value (Config.String _) = equals |-- Args.name >> Config.String; | 
| 24003 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 463 | |
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 464 | 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 | 465 | let val config_type = Config.get_global thy config | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 466 | 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 | 467 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 468 | fun register binding config thy = | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 469 | let val name = Sign.full_name thy binding in | 
| 24126 
913e1fa904fb
turned simp_depth_limit into configuration option;
 wenzelm parents: 
24110diff
changeset | 470 | thy | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 471 | |> 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 | 472 | |> Configs.map (Symtab.update (name, config)) | 
| 
913e1fa904fb
turned simp_depth_limit into configuration option;
 wenzelm parents: 
24110diff
changeset | 473 | end; | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 474 | |
| 42808 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 475 | fun declare make coerce binding default = | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 476 | let | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 477 | val name = Binding.name_of binding; | 
| 42808 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 478 |     val config_value = Config.declare_generic {global = false} name (make o default);
 | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 479 | val config = coerce config_value; | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 480 | 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 | 481 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 482 | in | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 483 | |
| 42808 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 484 | val config_bool = declare Config.Bool Config.bool; | 
| 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 485 | val config_int = declare Config.Int Config.int; | 
| 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 486 | val config_real = declare Config.Real Config.real; | 
| 
30870aee8a3f
discontinued global config options within attribute name space;
 wenzelm parents: 
42669diff
changeset | 487 | 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 | 488 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 489 | fun register_config config = register (Binding.name (Config.name_of config)) config; | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 490 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 491 | end; | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 492 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 493 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 494 | (* implicit setup *) | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 495 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 496 | local | 
| 24110 
4ab3084e311c
tuned config options: eliminated separate attribute "option";
 wenzelm parents: 
24030diff
changeset | 497 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 498 | 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 | 499 | let | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 500 | val (config, setup) = declare_config binding default; | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 501 | val _ = Context.>> (Context.map_theory setup); | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 502 | in config end; | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 503 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 504 | in | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 505 | |
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42380diff
changeset | 506 | 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 | 507 | 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 | 508 | 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 | 509 | val setup_config_real = setup_config config_real; | 
| 24003 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 510 | |
| 
da76d7e6435c
attribute "option": more elaborate syntax (with value parsing);
 wenzelm parents: 
23988diff
changeset | 511 | end; | 
| 23988 
aa46577f4f44
added attribute "option" for setting configuration options;
 wenzelm parents: 
23937diff
changeset | 512 | |
| 
aa46577f4f44
added attribute "option" for setting configuration options;
 wenzelm parents: 
23937diff
changeset | 513 | |
| 18636 | 514 | (* theory setup *) | 
| 5823 | 515 | |
| 26463 | 516 | val _ = Context.>> (Context.map_theory | 
| 42277 | 517 | (register_config Ast.trace_raw #> | 
| 43347 | 518 | register_config Ast.stats_raw #> | 
| 42289 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 519 | register_config Printer.show_brackets_raw #> | 
| 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 520 | register_config Printer.show_sorts_raw #> | 
| 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 521 | register_config Printer.show_types_raw #> | 
| 49657 
40e4feac2921
turn constraints into Isabelle_Markup.typing, depending on show_markup options;
 wenzelm parents: 
48992diff
changeset | 522 | register_config Printer.show_markup_raw #> | 
| 42289 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 523 | register_config Printer.show_structs_raw #> | 
| 
dafae095d733
discontinued special status of structure Printer;
 wenzelm parents: 
42284diff
changeset | 524 | register_config Printer.show_question_marks_raw #> | 
| 46512 
4f9f61f9b535
simplified configuration options for syntax ambiguity;
 wenzelm parents: 
46506diff
changeset | 525 | register_config Syntax.ambiguity_warning_raw #> | 
| 46506 
c7faa011bfa7
simplified configuration options for syntax ambiguity;
 wenzelm parents: 
46126diff
changeset | 526 | register_config Syntax.ambiguity_limit_raw #> | 
| 42284 | 527 | register_config Syntax_Trans.eta_contract_raw #> | 
| 42669 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 wenzelm parents: 
42616diff
changeset | 528 | register_config Name_Space.names_long_raw #> | 
| 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 wenzelm parents: 
42616diff
changeset | 529 | register_config Name_Space.names_short_raw #> | 
| 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 wenzelm parents: 
42616diff
changeset | 530 | register_config Name_Space.names_unique_raw #> | 
| 41375 | 531 | register_config ML_Context.trace_raw #> | 
| 42360 | 532 | 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 | 533 | 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 | 534 | 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 | 535 | register_config Goal_Display.show_consts_raw #> | 
| 39166 
19efc2af3e6c
turned show_hyps and show_tags into proper configuration option;
 wenzelm parents: 
39163diff
changeset | 536 | register_config Display.show_hyps_raw #> | 
| 
19efc2af3e6c
turned show_hyps and show_tags into proper configuration option;
 wenzelm parents: 
39163diff
changeset | 537 | register_config Display.show_tags_raw #> | 
| 39163 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 wenzelm parents: 
39137diff
changeset | 538 | 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 | 539 | 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 | 540 | 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 | 541 | register_config Unify.trace_types_raw #> | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41183diff
changeset | 542 | register_config Raw_Simplifier.simp_depth_limit_raw #> | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41183diff
changeset | 543 | register_config Raw_Simplifier.simp_trace_depth_limit_raw #> | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41183diff
changeset | 544 | register_config Raw_Simplifier.simp_debug_raw #> | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41183diff
changeset | 545 | register_config Raw_Simplifier.simp_trace_raw)); | 
| 5823 | 546 | |
| 547 | end; |