| author | wenzelm | 
| Sat, 10 Sep 2011 23:27:32 +0200 | |
| changeset 44872 | a98ef45122f3 | 
| parent 40291 | 012ed4426fda | 
| child 47814 | 53668571d300 | 
| 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  | 
|
| 
39163
 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 
wenzelm 
parents: 
39116 
diff
changeset
 | 
27  | 
  val declare_generic: {global: bool} -> string -> (Context.generic -> value) -> raw
 | 
| 
 
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  | 
| 
 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 
wenzelm 
parents: 
39116 
diff
changeset
 | 
29  | 
val declare: string -> (Context.generic -> value) -> raw  | 
| 24125 | 30  | 
val name_of: 'a T -> string  | 
| 24114 | 31  | 
end;  | 
32  | 
||
33  | 
structure Config: CONFIG =  | 
|
34  | 
struct  | 
|
35  | 
||
36  | 
(* simple values *)  | 
|
37  | 
||
38  | 
datatype value =  | 
|
39  | 
Bool of bool |  | 
|
40  | 
Int of int |  | 
|
| 40291 | 41  | 
Real of real |  | 
| 24114 | 42  | 
String of string;  | 
43  | 
||
44  | 
fun print_value (Bool true) = "true"  | 
|
45  | 
| print_value (Bool false) = "false"  | 
|
46  | 
| print_value (Int i) = signed_string_of_int i  | 
|
| 40291 | 47  | 
| print_value (Real x) = signed_string_of_real x  | 
| 24114 | 48  | 
| print_value (String s) = quote s;  | 
49  | 
||
| 38804 | 50  | 
fun print_type (Bool _) = "bool"  | 
51  | 
| print_type (Int _) = "int"  | 
|
| 40291 | 52  | 
| print_type (Real _) = "real"  | 
| 24114 | 53  | 
| print_type (String _) = "string";  | 
54  | 
||
55  | 
fun same_type (Bool _) (Bool _) = true  | 
|
56  | 
| same_type (Int _) (Int _) = true  | 
|
| 40291 | 57  | 
| same_type (Real _) (Real _) = true  | 
| 24114 | 58  | 
| same_type (String _) (String _) = true  | 
59  | 
| same_type _ _ = false;  | 
|
60  | 
||
61  | 
fun type_check name f value =  | 
|
62  | 
let  | 
|
63  | 
val value' = f value;  | 
|
64  | 
val _ = same_type value value' orelse  | 
|
65  | 
      error ("Ill-typed configuration option " ^ quote name ^ ": " ^
 | 
|
66  | 
print_type value ^ " expected,\nbut " ^ print_type value' ^ " was found");  | 
|
67  | 
in value' end;  | 
|
68  | 
||
69  | 
||
70  | 
(* abstract configuration options *)  | 
|
71  | 
||
72  | 
datatype 'a T = Config of  | 
|
| 24125 | 73  | 
 {name: string,
 | 
74  | 
get_value: Context.generic -> 'a,  | 
|
| 24114 | 75  | 
  map_value: ('a -> 'a) -> Context.generic -> Context.generic};
 | 
76  | 
||
| 
39163
 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 
wenzelm 
parents: 
39116 
diff
changeset
 | 
77  | 
type raw = value T;  | 
| 
 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 
wenzelm 
parents: 
39116 
diff
changeset
 | 
78  | 
|
| 24125 | 79  | 
fun coerce make dest (Config {name, get_value, map_value}) = Config
 | 
80  | 
 {name = name,
 | 
|
81  | 
get_value = dest o get_value,  | 
|
82  | 
map_value = fn f => map_value (make o f o dest)};  | 
|
| 24114 | 83  | 
|
84  | 
val bool = coerce Bool (fn Bool b => b);  | 
|
85  | 
val int = coerce Int (fn Int i => i);  | 
|
| 40291 | 86  | 
val real = coerce Real (fn Real x => x);  | 
| 24114 | 87  | 
val string = coerce String (fn String s => s);  | 
88  | 
||
89  | 
fun get_generic context (Config {get_value, ...}) = get_value context;
 | 
|
90  | 
fun map_generic (Config {map_value, ...}) f context = map_value f context;
 | 
|
91  | 
fun put_generic config value = map_generic config (K value);  | 
|
92  | 
||
93  | 
fun get_ctxt ctxt = get_generic (Context.Proof ctxt);  | 
|
94  | 
fun map_ctxt config f = Context.proof_map (map_generic config f);  | 
|
95  | 
fun put_ctxt config value = map_ctxt config (K value);  | 
|
96  | 
||
| 
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
 | 
97  | 
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
 | 
98  | 
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
 | 
99  | 
fun put_global config value = map_global config (K value);  | 
| 24114 | 100  | 
|
101  | 
||
102  | 
(* context information *)  | 
|
103  | 
||
| 33519 | 104  | 
structure Value = Generic_Data  | 
| 24114 | 105  | 
(  | 
106  | 
type T = value Inttab.table;  | 
|
107  | 
val empty = Inttab.empty;  | 
|
108  | 
val extend = I;  | 
|
| 33519 | 109  | 
fun merge data = Inttab.merge (K true) data;  | 
| 24114 | 110  | 
);  | 
111  | 
||
| 
39116
 
f14735a88886
more explicit Config.declare vs. Config.declare_global;
 
wenzelm 
parents: 
38804 
diff
changeset
 | 
112  | 
fun declare_generic {global} name default =
 | 
| 24114 | 113  | 
let  | 
114  | 
val id = serial ();  | 
|
115  | 
||
| 
36000
 
5560b2437789
configuration options admit dynamic default values;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
116  | 
fun get_value context =  | 
| 
 
5560b2437789
configuration options admit dynamic default values;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
117  | 
(case Inttab.lookup (Value.get context) id of  | 
| 
 
5560b2437789
configuration options admit dynamic default values;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
118  | 
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
 | 
119  | 
| NONE => default context);  | 
| 
36000
 
5560b2437789
configuration options admit dynamic default values;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
120  | 
|
| 
 
5560b2437789
configuration options admit dynamic default values;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
121  | 
fun update_value f context =  | 
| 
 
5560b2437789
configuration options admit dynamic default values;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
122  | 
Value.map (Inttab.update (id, type_check name f (get_value context))) context;  | 
| 24114 | 123  | 
|
124  | 
fun map_value f (context as Context.Proof _) =  | 
|
| 
36000
 
5560b2437789
configuration options admit dynamic default values;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
125  | 
let val context' = update_value f context in  | 
| 24114 | 126  | 
if global andalso  | 
| 40291 | 127  | 
print_value (get_value (Context.Theory (Context.theory_of context'))) <>  | 
128  | 
print_value (get_value context')  | 
|
| 24114 | 129  | 
            then (warning ("Ignoring local change of global option " ^ quote name); context)
 | 
130  | 
else context'  | 
|
131  | 
end  | 
|
| 
36000
 
5560b2437789
configuration options admit dynamic default values;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
132  | 
| map_value f context = update_value f context;  | 
| 24125 | 133  | 
  in Config {name = name, get_value = get_value, map_value = map_value} end;
 | 
134  | 
||
| 
39116
 
f14735a88886
more explicit Config.declare vs. Config.declare_global;
 
wenzelm 
parents: 
38804 
diff
changeset
 | 
135  | 
val declare_global = declare_generic {global = true};
 | 
| 
 
f14735a88886
more explicit Config.declare vs. Config.declare_global;
 
wenzelm 
parents: 
38804 
diff
changeset
 | 
136  | 
val declare = declare_generic {global = false};
 | 
| 
 
f14735a88886
more explicit Config.declare vs. Config.declare_global;
 
wenzelm 
parents: 
38804 
diff
changeset
 | 
137  | 
|
| 24125 | 138  | 
fun name_of (Config {name, ...}) = name;
 | 
| 24114 | 139  | 
|
140  | 
||
141  | 
(*final declarations of this structure!*)  | 
|
142  | 
val get = get_ctxt;  | 
|
143  | 
val map = map_ctxt;  | 
|
144  | 
val put = put_ctxt;  | 
|
145  | 
||
146  | 
end;  |