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