author | Christian Urban <urbanc@in.tum.de> |
Wed, 11 Mar 2009 22:34:49 +0000 | |
changeset 168 | 009ca4807baa |
parent 119 | 4536782969fa |
permissions | -rw-r--r-- |
61 | 1 |
theory Config |
63 | 2 |
imports "../Base" |
61 | 3 |
begin |
4 |
||
119
4536782969fa
added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents:
102
diff
changeset
|
5 |
section {* Configuration Options\label{rec:config} *} |
61 | 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 |
||
168
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
119
diff
changeset
|
15 |
Assume you want to control three values, say @{text bval} containing a |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
74
diff
changeset
|
16 |
boolean, @{text ival} containing an integer and @{text sval} |
168
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
119
diff
changeset
|
17 |
containing a string. These values can be declared on the ML-level by |
61 | 18 |
*} |
19 |
||
69
19106a9975c1
highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents:
68
diff
changeset
|
20 |
ML{*val (bval, setup_bval) = Attrib.config_bool "bval" false |
63 | 21 |
val (ival, setup_ival) = Attrib.config_int "ival" 0 |
69
19106a9975c1
highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents:
68
diff
changeset
|
22 |
val (sval, setup_sval) = Attrib.config_string "sval" "some string" *} |
63 | 23 |
|
24 |
text {* |
|
67 | 25 |
where each value needs to be given a default. To enable these values, they need to |
168
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
119
diff
changeset
|
26 |
be set up with |
63 | 27 |
*} |
28 |
||
29 |
setup {* setup_bval *} |
|
30 |
setup {* setup_ival *} |
|
31 |
||
32 |
text {* or on the ML-level *} |
|
33 |
||
69
19106a9975c1
highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents:
68
diff
changeset
|
34 |
ML{*setup_sval @{theory} *} |
63 | 35 |
|
36 |
text {* |
|
67 | 37 |
The user can now manipulate the values from within Isabelle with the command |
61 | 38 |
*} |
39 |
||
63 | 40 |
declare [[bval = true, ival = 3]] |
41 |
||
42 |
text {* |
|
67 | 43 |
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
|
44 |
function @{ML Config.get}: |
63 | 45 |
|
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents:
69
diff
changeset
|
46 |
@{ML_response [display,gray] "Config.get @{context} bval" "true"} |
61 | 47 |
|
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents:
69
diff
changeset
|
48 |
@{ML_response [display,gray] "Config.get @{context} ival" "3"} |
63 | 49 |
|
50 |
The function @{ML Config.put} manipulates the values. For example |
|
51 |
||
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents:
69
diff
changeset
|
52 |
@{ML_response [display,gray] "Config.put sval \"foo\" @{context}; Config.get @{context} sval" "foo"} |
63 | 53 |
|
168
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
119
diff
changeset
|
54 |
The same can be achieved using the command \isacommand{setup}. |
61 | 55 |
*} |
56 |
||
63 | 57 |
setup {* Config.put_thy sval "bar" *} |
58 |
||
59 |
text {* |
|
168
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
119
diff
changeset
|
60 |
Now the retrival of this value yields: |
63 | 61 |
|
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents:
69
diff
changeset
|
62 |
@{ML_response [display,gray] "Config.get @{context} sval" "\"bar\""} |
63 | 63 |
|
64 |
We can apply a function to a value using @{ML Config.map}. For example incrementing |
|
168
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
119
diff
changeset
|
65 |
@{ML ival} can be done by: |
63 | 66 |
|
72
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents:
69
diff
changeset
|
67 |
@{ML_response [display,gray] |
63 | 68 |
"let |
69 |
val ctxt = Config.map ival (fn i => i + 1) @{context} |
|
70 |
in |
|
71 |
Config.get ctxt ival |
|
72 |
end" "4"} |
|
73 |
||
68
e7519207c2b7
added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents:
67
diff
changeset
|
74 |
\begin{readmore} |
63 | 75 |
For more information see @{ML_file "Pure/Isar/attrib.ML"} and @{ML_file "Pure/config.ML"}. |
68
e7519207c2b7
added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents:
67
diff
changeset
|
76 |
\end{readmore} |
63 | 77 |
|
68
e7519207c2b7
added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents:
67
diff
changeset
|
78 |
There are many good reasons to control parameters in this way. One is |
e7519207c2b7
added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents:
67
diff
changeset
|
79 |
that it avoid global references, which cause many headaches with the |
e7519207c2b7
added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents:
67
diff
changeset
|
80 |
multithreaded execution of Isabelle. |
e7519207c2b7
added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents:
67
diff
changeset
|
81 |
|
61 | 82 |
*} |
83 |
end |