author | Christian Urban <urbanc@in.tum.de> |
Mon, 12 Jan 2009 16:49:15 +0000 | |
changeset 67 | 5fbeeac2901b |
parent 65 | c8e9a4f97916 |
child 68 | e7519207c2b7 |
permissions | -rw-r--r-- |
61 | 1 |
theory Config |
63 | 2 |
imports "../Base" |
61 | 3 |
begin |
4 |
||
5 |
section {* Configuration Options *} |
|
6 |
||
7 |
||
8 |
text {* |
|
9 |
{\bf Problem:} |
|
63 | 10 |
You would like to enhance your tool with options that can be changed |
11 |
by the user without having to resort to the ML-level.\smallskip |
|
61 | 12 |
|
13 |
{\bf Solution:} This can be achieved using configuration values.\smallskip |
|
14 |
||
63 | 15 |
Assume you want to control three values, say @{ML_text bval} containing a |
16 |
boolean, @{ML_text ival} containing an integer and @{ML_text sval} |
|
67 | 17 |
containing a string. These values can be declared on the ML-level with |
61 | 18 |
*} |
19 |
||
20 |
ML {* |
|
63 | 21 |
val (bval, setup_bval) = Attrib.config_bool "bval" false |
22 |
val (ival, setup_ival) = Attrib.config_int "ival" 0 |
|
23 |
val (sval, setup_sval) = Attrib.config_string "sval" "some string" |
|
24 |
*} |
|
25 |
||
26 |
text {* |
|
67 | 27 |
where each value needs to be given a default. To enable these values, they need to |
63 | 28 |
be set up using |
29 |
*} |
|
30 |
||
31 |
setup {* setup_bval *} |
|
32 |
setup {* setup_ival *} |
|
33 |
||
34 |
text {* or on the ML-level *} |
|
35 |
||
65
c8e9a4f97916
tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents:
63
diff
changeset
|
36 |
ML {* |
c8e9a4f97916
tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents:
63
diff
changeset
|
37 |
setup_sval @{theory} |
63 | 38 |
*} |
39 |
||
40 |
text {* |
|
67 | 41 |
The user can now manipulate the values from within Isabelle with the command |
61 | 42 |
*} |
43 |
||
63 | 44 |
declare [[bval = true, ival = 3]] |
45 |
||
46 |
text {* |
|
67 | 47 |
On the ML-level these values can be retrieved using the |
65
c8e9a4f97916
tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents:
63
diff
changeset
|
48 |
function @{ML Config.get}: |
63 | 49 |
|
50 |
@{ML_response [display] "Config.get @{context} bval" "true"} |
|
61 | 51 |
|
63 | 52 |
@{ML_response [display] "Config.get @{context} ival" "3"} |
53 |
||
54 |
The function @{ML Config.put} manipulates the values. For example |
|
55 |
||
56 |
@{ML_response [display] "Config.put sval \"foo\" @{context}; Config.get @{context} sval" "foo"} |
|
57 |
||
58 |
The same can be achived using the command \isacommand{setup}. |
|
61 | 59 |
*} |
60 |
||
63 | 61 |
setup {* Config.put_thy sval "bar" *} |
62 |
||
63 |
text {* |
|
65
c8e9a4f97916
tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents:
63
diff
changeset
|
64 |
The retrival of this value yields now |
63 | 65 |
|
66 |
@{ML_response [display] "Config.get @{context} sval" "\"bar\""} |
|
67 |
||
68 |
We can apply a function to a value using @{ML Config.map}. For example incrementing |
|
69 |
@{ML ival} can be done by |
|
70 |
||
71 |
@{ML_response [display] |
|
72 |
"let |
|
73 |
val ctxt = Config.map ival (fn i => i + 1) @{context} |
|
74 |
in |
|
75 |
Config.get ctxt ival |
|
76 |
end" "4"} |
|
77 |
||
78 |
\begin{readmore} |
|
79 |
For more information see @{ML_file "Pure/Isar/attrib.ML"} and @{ML_file "Pure/config.ML"}. |
|
80 |
\end{readmore} |
|
81 |
||
61 | 82 |
*} |
83 |
||
84 |
||
85 |
text {* |
|
86 |
||
87 |
Note: Avoid to use references for this purpose! |
|
88 |
*} |
|
89 |
||
90 |
end |