author | wenzelm |
Wed, 28 Aug 2013 22:50:23 +0200 | |
changeset 53252 | 4766fbe322b5 |
parent 52469 | c06f1d36a8c9 |
child 56294 | 85911b8a6868 |
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 |
24114 | 10 |
val print_value: value -> string |
11 |
val print_type: value -> string |
|
12 |
type 'a T |
|
39163
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
wenzelm
parents:
39116
diff
changeset
|
13 |
type raw = value T |
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
wenzelm
parents:
39116
diff
changeset
|
14 |
val bool: raw -> bool T |
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
wenzelm
parents:
39116
diff
changeset
|
15 |
val int: raw -> int T |
40291 | 16 |
val real: raw -> real T |
39163
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
wenzelm
parents:
39116
diff
changeset
|
17 |
val string: raw -> string T |
24114 | 18 |
val get: Proof.context -> 'a T -> 'a |
19 |
val map: 'a T -> ('a -> 'a) -> Proof.context -> Proof.context |
|
20 |
val put: 'a T -> 'a -> Proof.context -> Proof.context |
|
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
|
21 |
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
|
22 |
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
|
23 |
val put_global: 'a T -> 'a -> theory -> theory |
24114 | 24 |
val get_generic: Context.generic -> 'a T -> 'a |
25 |
val map_generic: 'a T -> ('a -> 'a) -> Context.generic -> Context.generic |
|
26 |
val put_generic: 'a T -> 'a -> Context.generic -> Context.generic |
|
52469 | 27 |
val declare: string -> (Context.generic -> value) -> raw |
39163
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
wenzelm
parents:
39116
diff
changeset
|
28 |
val declare_global: string -> (Context.generic -> value) -> raw |
51947
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
29 |
val declare_option: string -> raw |
52469 | 30 |
val declare_option_global: string -> raw |
24125 | 31 |
val name_of: 'a T -> string |
24114 | 32 |
end; |
33 |
||
34 |
structure Config: CONFIG = |
|
35 |
struct |
|
36 |
||
37 |
(* simple values *) |
|
38 |
||
39 |
datatype value = |
|
40 |
Bool of bool | |
|
41 |
Int of int | |
|
40291 | 42 |
Real of real | |
24114 | 43 |
String of string; |
44 |
||
45 |
fun print_value (Bool true) = "true" |
|
46 |
| print_value (Bool false) = "false" |
|
47 |
| print_value (Int i) = signed_string_of_int i |
|
51990 | 48 |
| print_value (Real x) = Markup.print_real x |
24114 | 49 |
| print_value (String s) = quote s; |
50 |
||
38804 | 51 |
fun print_type (Bool _) = "bool" |
52 |
| print_type (Int _) = "int" |
|
40291 | 53 |
| print_type (Real _) = "real" |
24114 | 54 |
| print_type (String _) = "string"; |
55 |
||
56 |
fun same_type (Bool _) (Bool _) = true |
|
57 |
| same_type (Int _) (Int _) = true |
|
40291 | 58 |
| same_type (Real _) (Real _) = true |
24114 | 59 |
| same_type (String _) (String _) = true |
60 |
| same_type _ _ = false; |
|
61 |
||
62 |
fun type_check name f value = |
|
63 |
let |
|
64 |
val value' = f value; |
|
65 |
val _ = same_type value value' orelse |
|
66 |
error ("Ill-typed configuration option " ^ quote name ^ ": " ^ |
|
67 |
print_type value ^ " expected,\nbut " ^ print_type value' ^ " was found"); |
|
68 |
in value' end; |
|
69 |
||
70 |
||
71 |
(* abstract configuration options *) |
|
72 |
||
73 |
datatype 'a T = Config of |
|
24125 | 74 |
{name: string, |
75 |
get_value: Context.generic -> 'a, |
|
24114 | 76 |
map_value: ('a -> 'a) -> Context.generic -> Context.generic}; |
77 |
||
39163
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
wenzelm
parents:
39116
diff
changeset
|
78 |
type raw = value T; |
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
wenzelm
parents:
39116
diff
changeset
|
79 |
|
24125 | 80 |
fun coerce make dest (Config {name, get_value, map_value}) = Config |
81 |
{name = name, |
|
82 |
get_value = dest o get_value, |
|
83 |
map_value = fn f => map_value (make o f o dest)}; |
|
24114 | 84 |
|
85 |
val bool = coerce Bool (fn Bool b => b); |
|
86 |
val int = coerce Int (fn Int i => i); |
|
40291 | 87 |
val real = coerce Real (fn Real x => x); |
24114 | 88 |
val string = coerce String (fn String s => s); |
89 |
||
90 |
fun get_generic context (Config {get_value, ...}) = get_value context; |
|
91 |
fun map_generic (Config {map_value, ...}) f context = map_value f context; |
|
92 |
fun put_generic config value = map_generic config (K value); |
|
93 |
||
94 |
fun get_ctxt ctxt = get_generic (Context.Proof ctxt); |
|
95 |
fun map_ctxt config f = Context.proof_map (map_generic config f); |
|
96 |
fun put_ctxt config value = map_ctxt config (K value); |
|
97 |
||
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
|
98 |
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
|
99 |
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
|
100 |
fun put_global config value = map_global config (K value); |
24114 | 101 |
|
102 |
||
103 |
(* context information *) |
|
104 |
||
33519 | 105 |
structure Value = Generic_Data |
24114 | 106 |
( |
107 |
type T = value Inttab.table; |
|
108 |
val empty = Inttab.empty; |
|
109 |
val extend = I; |
|
33519 | 110 |
fun merge data = Inttab.merge (K true) data; |
24114 | 111 |
); |
112 |
||
52469 | 113 |
local |
114 |
||
52039 | 115 |
fun declare_generic global name default = |
24114 | 116 |
let |
117 |
val id = serial (); |
|
118 |
||
36000
5560b2437789
configuration options admit dynamic default values;
wenzelm
parents:
33519
diff
changeset
|
119 |
fun get_value context = |
5560b2437789
configuration options admit dynamic default values;
wenzelm
parents:
33519
diff
changeset
|
120 |
(case Inttab.lookup (Value.get context) id of |
5560b2437789
configuration options admit dynamic default values;
wenzelm
parents:
33519
diff
changeset
|
121 |
SOME value => value |
36002
f4f343500249
pass raw Context.generic, to avoid wasteful Context.proof_of -- Config.get_thy is often used in performance critical spots like unify.ML;
wenzelm
parents:
36000
diff
changeset
|
122 |
| NONE => default context); |
36000
5560b2437789
configuration options admit dynamic default values;
wenzelm
parents:
33519
diff
changeset
|
123 |
|
5560b2437789
configuration options admit dynamic default values;
wenzelm
parents:
33519
diff
changeset
|
124 |
fun update_value f context = |
5560b2437789
configuration options admit dynamic default values;
wenzelm
parents:
33519
diff
changeset
|
125 |
Value.map (Inttab.update (id, type_check name f (get_value context))) context; |
24114 | 126 |
|
47814
53668571d300
avoid spurious warning in invisible context, notably Haftmann-Wenzel sandwich;
wenzelm
parents:
40291
diff
changeset
|
127 |
fun map_value f (context as Context.Proof ctxt) = |
36000
5560b2437789
configuration options admit dynamic default values;
wenzelm
parents:
33519
diff
changeset
|
128 |
let val context' = update_value f context in |
24114 | 129 |
if global andalso |
47814
53668571d300
avoid spurious warning in invisible context, notably Haftmann-Wenzel sandwich;
wenzelm
parents:
40291
diff
changeset
|
130 |
Context_Position.is_visible ctxt andalso |
40291 | 131 |
print_value (get_value (Context.Theory (Context.theory_of context'))) <> |
132 |
print_value (get_value context') |
|
47814
53668571d300
avoid spurious warning in invisible context, notably Haftmann-Wenzel sandwich;
wenzelm
parents:
40291
diff
changeset
|
133 |
then |
53668571d300
avoid spurious warning in invisible context, notably Haftmann-Wenzel sandwich;
wenzelm
parents:
40291
diff
changeset
|
134 |
(Context_Position.if_visible ctxt warning |
53668571d300
avoid spurious warning in invisible context, notably Haftmann-Wenzel sandwich;
wenzelm
parents:
40291
diff
changeset
|
135 |
("Ignoring local change of global option " ^ quote name); context) |
24114 | 136 |
else context' |
137 |
end |
|
36000
5560b2437789
configuration options admit dynamic default values;
wenzelm
parents:
33519
diff
changeset
|
138 |
| map_value f context = update_value f context; |
24125 | 139 |
in Config {name = name, get_value = get_value, map_value = map_value} end; |
140 |
||
52469 | 141 |
fun declare_option_generic global name = |
51947
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
142 |
let |
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
143 |
val typ = Options.default_typ name; |
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
144 |
val default = |
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
145 |
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
|
146 |
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
|
147 |
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
|
148 |
else if typ = Options.stringT then fn _ => String (Options.default_string name) |
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
149 |
else error ("Unknown type for option " ^ quote name ^ " : " ^ quote typ); |
52469 | 150 |
in declare_generic global name default end; |
151 |
||
152 |
in |
|
153 |
||
154 |
val declare = declare_generic false; |
|
155 |
val declare_global = declare_generic true; |
|
156 |
val declare_option = declare_option_generic false; |
|
157 |
val declare_option_global = declare_option_generic true; |
|
158 |
||
159 |
end; |
|
51947
3301612c4893
support for system options as context-sensitive config options;
wenzelm
parents:
47814
diff
changeset
|
160 |
|
24125 | 161 |
fun name_of (Config {name, ...}) = name; |
24114 | 162 |
|
163 |
||
164 |
(*final declarations of this structure!*) |
|
165 |
val get = get_ctxt; |
|
166 |
val map = map_ctxt; |
|
167 |
val put = put_ctxt; |
|
168 |
||
169 |
end; |