author | wenzelm |
Thu, 10 Apr 2025 14:12:33 +0200 | |
changeset 82475 | 0a6d57c4d58b |
parent 80901 | f4d519d088af |
permissions | -rw-r--r-- |
24114 | 1 |
(* Title: Pure/config.ML |
2 |
Author: Makarius |
|
3 |
||
4 |
Configuration options as values within the local context. |
|
5 |
*) |
|
6 |
||
7 |
signature CONFIG = |
|
8 |
sig |
|
40291 | 9 |
datatype value = Bool of bool | Int of int | Real of real | String of string |
80901
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
10 |
val eq_value: value * value -> bool |
24114 | 11 |
val print_value: value -> string |
12 |
val print_type: value -> string |
|
13 |
type 'a T |
|
69575 | 14 |
val name_of: 'a T -> string |
15 |
val pos_of: 'a T -> Position.T |
|
69576 | 16 |
val apply: 'a T -> Proof.context -> 'a |
24114 | 17 |
val get: Proof.context -> 'a T -> 'a |
18 |
val map: 'a T -> ('a -> 'a) -> Proof.context -> Proof.context |
|
19 |
val put: 'a T -> 'a -> Proof.context -> Proof.context |
|
68827 | 20 |
val restore: 'a T -> Proof.context -> Proof.context -> Proof.context |
69576 | 21 |
val apply_global: 'a T -> theory -> 'a |
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
|
22 |
val get_global: theory -> 'a T -> 'a |
f60e4dd6d76f
renamed Config.get_thy to Config.get_global etc. to indicate that this is not the real thing;
wenzelm
parents:
36002
diff
changeset
|
23 |
val map_global: 'a T -> ('a -> 'a) -> theory -> theory |
f60e4dd6d76f
renamed Config.get_thy to Config.get_global etc. to indicate that this is not the real thing;
wenzelm
parents:
36002
diff
changeset
|
24 |
val put_global: 'a T -> 'a -> theory -> theory |
68827 | 25 |
val restore_global: 'a T -> theory -> theory -> theory |
69576 | 26 |
val apply_generic: 'a T -> Context.generic -> 'a |
24114 | 27 |
val get_generic: Context.generic -> 'a T -> 'a |
28 |
val map_generic: 'a T -> ('a -> 'a) -> Context.generic -> Context.generic |
|
29 |
val put_generic: 'a T -> 'a -> Context.generic -> Context.generic |
|
68827 | 30 |
val restore_generic: 'a T -> Context.generic -> Context.generic -> Context.generic |
69575 | 31 |
val bool: value T -> bool T |
32 |
val bool_value: bool T -> value T |
|
33 |
val int: value T -> int T |
|
34 |
val int_value: int T -> value T |
|
35 |
val real: value T -> real T |
|
36 |
val real_value: real T -> value T |
|
37 |
val string: value T -> string T |
|
38 |
val string_value: string T -> value T |
|
39 |
val declare: string * Position.T -> (Context.generic -> value) -> value T |
|
40 |
val declare_bool: string * Position.T -> (Context.generic -> bool) -> bool T |
|
41 |
val declare_int: string * Position.T -> (Context.generic -> int) -> int T |
|
42 |
val declare_real: string * Position.T -> (Context.generic -> real) -> real T |
|
43 |
val declare_string: string * Position.T -> (Context.generic -> string) -> string T |
|
44 |
val declare_option: string * Position.T -> value T |
|
45 |
val declare_option_bool: string * Position.T -> bool T |
|
46 |
val declare_option_int: string * Position.T -> int T |
|
47 |
val declare_option_real: string * Position.T -> real T |
|
48 |
val declare_option_string: string * Position.T -> string T |
|
24114 | 49 |
end; |
50 |
||
51 |
structure Config: CONFIG = |
|
52 |
struct |
|
53 |
||
54 |
(* simple values *) |
|
55 |
||
56 |
datatype value = |
|
57 |
Bool of bool | |
|
58 |
Int of int | |
|
40291 | 59 |
Real of real | |
24114 | 60 |
String of string; |
61 |
||
80901
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
62 |
fun eq_value (Bool a, Bool b) = a = b |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
63 |
| eq_value (Int a, Int b) = a = b |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
64 |
| eq_value (Real a, Real b) = eq_real (a, b) |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
65 |
| eq_value (String a, String b) = a = b |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
66 |
| eq_value _ = false; |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
67 |
|
24114 | 68 |
fun print_value (Bool true) = "true" |
69 |
| print_value (Bool false) = "false" |
|
69574 | 70 |
| print_value (Int i) = Value.print_int i |
63806 | 71 |
| print_value (Real x) = Value.print_real x |
24114 | 72 |
| print_value (String s) = quote s; |
73 |
||
38804 | 74 |
fun print_type (Bool _) = "bool" |
75 |
| print_type (Int _) = "int" |
|
40291 | 76 |
| print_type (Real _) = "real" |
24114 | 77 |
| print_type (String _) = "string"; |
78 |
||
79 |
fun same_type (Bool _) (Bool _) = true |
|
80 |
| same_type (Int _) (Int _) = true |
|
40291 | 81 |
| same_type (Real _) (Real _) = true |
24114 | 82 |
| same_type (String _) (String _) = true |
83 |
| same_type _ _ = false; |
|
84 |
||
56438 | 85 |
fun type_check (name, pos) f value = |
24114 | 86 |
let |
87 |
val value' = f value; |
|
88 |
val _ = same_type value value' orelse |
|
56438 | 89 |
error ("Ill-typed configuration option " ^ quote name ^ Position.here pos ^ ": " ^ |
24114 | 90 |
print_type value ^ " expected,\nbut " ^ print_type value' ^ " was found"); |
91 |
in value' end; |
|
92 |
||
93 |
||
94 |
(* abstract configuration options *) |
|
95 |
||
96 |
datatype 'a T = Config of |
|
24125 | 97 |
{name: string, |
56438 | 98 |
pos: Position.T, |
24125 | 99 |
get_value: Context.generic -> 'a, |
24114 | 100 |
map_value: ('a -> 'a) -> Context.generic -> Context.generic}; |
101 |
||
69575 | 102 |
fun name_of (Config {name, ...}) = name; |
103 |
fun pos_of (Config {pos, ...}) = pos; |
|
24114 | 104 |
|
69576 | 105 |
fun apply_generic (Config {get_value, ...}) = get_value; |
106 |
fun get_generic context config = apply_generic config context; |
|
24114 | 107 |
fun map_generic (Config {map_value, ...}) f context = map_value f context; |
108 |
fun put_generic config value = map_generic config (K value); |
|
69576 | 109 |
fun restore_generic config = put_generic config o apply_generic config; |
24114 | 110 |
|
69576 | 111 |
fun apply_ctxt config = apply_generic config o Context.Proof; |
24114 | 112 |
fun get_ctxt ctxt = get_generic (Context.Proof ctxt); |
113 |
fun map_ctxt config f = Context.proof_map (map_generic config f); |
|
114 |
fun put_ctxt config value = map_ctxt config (K value); |
|
69576 | 115 |
fun restore_ctxt config = put_ctxt config o apply_ctxt config; |
24114 | 116 |
|
69576 | 117 |
fun apply_global config = apply_generic config o Context.Theory; |
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
|
118 |
fun get_global thy = get_generic (Context.Theory thy); |
f60e4dd6d76f
renamed Config.get_thy to Config.get_global etc. to indicate that this is not the real thing;
wenzelm
parents:
36002
diff
changeset
|
119 |
fun map_global config f = Context.theory_map (map_generic config f); |
f60e4dd6d76f
renamed Config.get_thy to Config.get_global etc. to indicate that this is not the real thing;
wenzelm
parents:
36002
diff
changeset
|
120 |
fun put_global config value = map_global config (K value); |
69576 | 121 |
fun restore_global config = put_global config o apply_global config; |
24114 | 122 |
|
123 |
||
69575 | 124 |
(* coerce type *) |
125 |
||
126 |
fun coerce make dest (Config {name, pos, get_value, map_value}) = Config |
|
127 |
{name = name, |
|
128 |
pos = pos, |
|
129 |
get_value = dest o get_value, |
|
130 |
map_value = fn f => map_value (make o f o dest)}; |
|
131 |
||
132 |
fun coerce_pair make dest = (coerce make dest, coerce dest make); |
|
133 |
||
134 |
val (bool, bool_value) = coerce_pair Bool (fn Bool b => b); |
|
135 |
val (int, int_value) = coerce_pair Int (fn Int i => i); |
|
136 |
val (real, real_value) = coerce_pair Real (fn Real x => x); |
|
137 |
val (string, string_value) = coerce_pair String (fn String s => s); |
|
138 |
||
139 |
||
140 |
(* context data *) |
|
24114 | 141 |
|
63806 | 142 |
structure Data = Generic_Data |
24114 | 143 |
( |
144 |
type T = value Inttab.table; |
|
145 |
val empty = Inttab.empty; |
|
33519 | 146 |
fun merge data = Inttab.merge (K true) data; |
24114 | 147 |
); |
148 |
||
58951
8b7caf447357
removed obsolete global-only options, which did not work out anyway (due to complexity of local_theory sandwich);
wenzelm
parents:
57858
diff
changeset
|
149 |
fun declare (name, pos) default = |
24114 | 150 |
let |
151 |
val id = serial (); |
|
152 |
||
80901
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
153 |
val lookup_value = Inttab.lookup o Data.get; |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
154 |
|
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
155 |
fun the_default_value _ (SOME value) = value |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
156 |
| the_default_value context NONE = default context; |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
157 |
|
36000
5560b2437789
configuration options admit dynamic default values;
wenzelm
parents:
33519
diff
changeset
|
158 |
fun get_value context = |
80901
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
159 |
the_default_value context (lookup_value context id); |
36000
5560b2437789
configuration options admit dynamic default values;
wenzelm
parents:
33519
diff
changeset
|
160 |
|
58951
8b7caf447357
removed obsolete global-only options, which did not work out anyway (due to complexity of local_theory sandwich);
wenzelm
parents:
57858
diff
changeset
|
161 |
fun map_value f context = |
80901
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
162 |
let |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
163 |
val old = lookup_value context id; |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
164 |
val old_value = the_default_value context old; |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
165 |
val new_value = type_check (name, pos) f old_value; |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
166 |
in |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
167 |
if eq_option eq_value (old, SOME new_value) then context |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
168 |
else Data.map (Inttab.update (id, new_value)) context |
f4d519d088af
minor performance tuning: avoid vacuous update of context;
wenzelm
parents:
76058
diff
changeset
|
169 |
end; |
58951
8b7caf447357
removed obsolete global-only options, which did not work out anyway (due to complexity of local_theory sandwich);
wenzelm
parents:
57858
diff
changeset
|
170 |
in |
8b7caf447357
removed obsolete global-only options, which did not work out anyway (due to complexity of local_theory sandwich);
wenzelm
parents:
57858
diff
changeset
|
171 |
Config {name = name, pos = pos, get_value = get_value, map_value = map_value} |
8b7caf447357
removed obsolete global-only options, which did not work out anyway (due to complexity of local_theory sandwich);
wenzelm
parents:
57858
diff
changeset
|
172 |
end; |
24114 | 173 |
|
69575 | 174 |
fun declare_bool name default = bool (declare name (Bool o default)); |
175 |
fun declare_int name default = int (declare name (Int o default)); |
|
176 |
fun declare_real name default = real (declare name (Real o default)); |
|
177 |
fun declare_string name default = string (declare name (String o default)); |
|
178 |
||
179 |
||
180 |
(* system options *) |
|
181 |
||
58951
8b7caf447357
removed obsolete global-only options, which did not work out anyway (due to complexity of local_theory sandwich);
wenzelm
parents:
57858
diff
changeset
|
182 |
fun declare_option (name, pos) = |
51947
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
183 |
let |
76058 | 184 |
val typ = Options.default_typ name handle ERROR msg => error (msg ^ Position.here pos); |
51947
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
185 |
val default = |
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
186 |
if typ = Options.boolT then fn _ => Bool (Options.default_bool name) |
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
187 |
else if typ = Options.intT then fn _ => Int (Options.default_int name) |
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
188 |
else if typ = Options.realT then fn _ => Real (Options.default_real name) |
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
189 |
else if typ = Options.stringT then fn _ => String (Options.default_string name) |
56438 | 190 |
else error ("Unknown type for option " ^ quote name ^ Position.here pos ^ " : " ^ quote typ); |
58951
8b7caf447357
removed obsolete global-only options, which did not work out anyway (due to complexity of local_theory sandwich);
wenzelm
parents:
57858
diff
changeset
|
191 |
in declare (name, pos) default end; |
51947
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
192 |
|
69575 | 193 |
val declare_option_bool = bool o declare_option; |
194 |
val declare_option_int = int o declare_option; |
|
195 |
val declare_option_real = real o declare_option; |
|
196 |
val declare_option_string = string o declare_option; |
|
24114 | 197 |
|
198 |
(*final declarations of this structure!*) |
|
69576 | 199 |
val apply = apply_ctxt; |
24114 | 200 |
val get = get_ctxt; |
201 |
val map = map_ctxt; |
|
202 |
val put = put_ctxt; |
|
68827 | 203 |
val restore = restore_ctxt; |
24114 | 204 |
|
205 |
end; |