| author | Christian Urban <urbanc@in.tum.de> |
| Fri, 13 Feb 2009 01:05:09 +0000 | |
| changeset 112 | a90d0fb24e75 |
| parent 102 | 5e309df58557 |
| child 119 | 4536782969fa |
| permissions | -rw-r--r-- |
| 61 | 1 |
theory ExternalSolver |
2 |
imports "../Base" |
|
| 94 | 3 |
uses ("external_solver.ML")
|
| 61 | 4 |
begin |
5 |
||
| 94 | 6 |
|
7 |
section {* Executing an External Application *}
|
|
| 61 | 8 |
|
9 |
text {*
|
|
10 |
{\bf Problem:}
|
|
| 94 | 11 |
You want to use an external application. |
12 |
\smallskip |
|
13 |
||
14 |
{\bf Solution:} The function @{ML system_out} might be the right thing for
|
|
15 |
you. |
|
| 61 | 16 |
\smallskip |
17 |
||
| 94 | 18 |
This function executes an external command as if printed in a shell. It |
19 |
returns the output of the program and its return value. |
|
20 |
||
21 |
For example, consider running an ordinary shell commands: |
|
22 |
||
23 |
@{ML_response [display,gray]
|
|
24 |
"system_out \"echo Hello world!\"" "(\"Hello world!\\n\", 0)"} |
|
25 |
||
26 |
Note that it works also fine with timeouts (see Recipe~\ref{rec:timeout}
|
|
27 |
on Page~\pageref{rec:timeout}), i.e. external applications are killed
|
|
28 |
properly. For example, the following expression takes only approximately |
|
29 |
one second: |
|
30 |
||
31 |
@{ML_response [display,gray]
|
|
32 |
"TimeLimit.timeLimit (Time.fromSeconds 1) system_out \"sleep 30\" |
|
33 |
handle TimeLimit.TimeOut => (\"timeout\", ~1)" "(\"timeout\", ~1)"} |
|
34 |
*} |
|
35 |
||
36 |
text {*
|
|
37 |
The function @{ML system_out} can also be used for more reasonable
|
|
38 |
applications, e.g. coupling external solvers with Isabelle. In that case, |
|
39 |
one has to make sure that Isabelle can find the particular executable. |
|
40 |
One way to ensure this is by adding a Bash-like variable binding into |
|
41 |
one of Isabelle's settings file (prefer the user settings file usually to |
|
42 |
be found at @{text "$HOME/.isabelle/etc/settings"}).
|
|
43 |
||
44 |
For example, assume you want to use the application @{text foo} which
|
|
45 |
is here supposed to be located at @{text "/usr/local/bin/"}.
|
|
46 |
The following line has to be added to one of Isabelle's settings file: |
|
47 |
||
48 |
@{text "FOO=/usr/local/bin/foo"}
|
|
49 |
||
50 |
In Isabelle, this application may now be executed by |
|
51 |
||
52 |
@{ML_response_fake [display,gray] "system_out \"$FOO\"" "\<dots>"}
|
|
53 |
*} |
|
54 |
||
55 |
||
56 |
||
57 |
||
58 |
||
59 |
||
60 |
section {* Writing an Oracle\label{rec:external} *}
|
|
61 |
||
62 |
text {*
|
|
63 |
{\bf Problem:}
|
|
64 |
You want to use a fast, new decision procedure not based one Isabelle's |
|
65 |
tactics, and you do not care whether it is sound. |
|
66 |
\smallskip |
|
67 |
||
68 |
{\bf Solution:} Isabelle provides the oracle mechanisms to bypass the
|
|
69 |
inference kernel. Note that theorems proven by an oracle carry a special |
|
70 |
mark to inform the user of their potential incorrectness. |
|
| 61 | 71 |
\smallskip |
72 |
||
73 |
\begin{readmore}
|
|
74 |
A short introduction to oracles can be found in [isar-ref: no suitable label |
|
| 94 | 75 |
for section 3.11]. A simple example, which we will slightly extend here, |
|
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
94
diff
changeset
|
76 |
is given in @{ML_file "FOL/ex/IffOracle.thy"}. The raw interface for adding
|
|
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
94
diff
changeset
|
77 |
oracles is @{ML add_oracle in Thm} in @{ML_file "Pure/thm.ML"}.
|
| 61 | 78 |
\end{readmore}
|
79 |
||
| 94 | 80 |
For our explanation here, we restrict ourselves to decide propositional |
81 |
formulae which consist only of equivalences between propositional variables, |
|
82 |
i.e. we want to decide whether @{term "P = (Q = P) = Q"} is a tautology.
|
|
83 |
||
84 |
Assume, that we have a decision procedure for such formulae, implemented |
|
85 |
in ML. Since we do not care how it works, we will use it here as an |
|
86 |
``external solver'': |
|
87 |
*} |
|
88 |
||
89 |
use "external_solver.ML" |
|
| 61 | 90 |
|
| 94 | 91 |
text {*
|
92 |
We do, however, know that the solver provides a function |
|
93 |
@{ML IffSolver.decide}.
|
|
94 |
It takes a string representation of a formula and returns either |
|
95 |
@{ML true} if the formula is a tautology or
|
|
96 |
@{ML false} otherwise. The input syntax is specified as follows:
|
|
97 |
||
98 |
formula $::=$ atom $\mid$ \verb|(| formula \verb|<=>| formula \verb|)| |
|
99 |
||
100 |
and all token are separated by at least one space. |
|
101 |
||
102 |
(FIXME: is there a better way for describing the syntax?) |
|
103 |
||
104 |
We will proceed in the following way. We start by translating a HOL formula |
|
105 |
into the string representation expected by the solver. The solver's result |
|
106 |
is then used to build an oracle, which we will subsequently use as a core |
|
107 |
for an Isar method to be able to apply the oracle in proving theorems. |
|
108 |
||
109 |
Let us start with the translation function from Isabelle propositions into |
|
110 |
the solver's string representation. To increase efficiency while building |
|
111 |
the string, we use functions from the @{text Buffer} module.
|
|
| 61 | 112 |
*} |
113 |
||
| 94 | 114 |
ML {*fun translate t =
|
115 |
let |
|
116 |
fun trans t = |
|
117 |
(case t of |
|
118 |
@{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t $ u =>
|
|
119 |
Buffer.add " (" #>
|
|
120 |
trans t #> |
|
121 |
Buffer.add "<=>" #> |
|
122 |
trans u #> |
|
123 |
Buffer.add ") " |
|
124 |
| Free (n, @{typ bool}) =>
|
|
125 |
Buffer.add " " #> |
|
126 |
Buffer.add n #> |
|
127 |
Buffer.add " " |
|
128 |
| _ => error "inacceptable term") |
|
129 |
in Buffer.content (trans t Buffer.empty) end |
|
| 61 | 130 |
*} |
131 |
||
| 94 | 132 |
text {*
|
133 |
Here is the string representation of the term @{term "p = (q = p)"}:
|
|
| 61 | 134 |
|
| 94 | 135 |
@{ML_response
|
136 |
"translate @{term \"p = (q = p)\"}"
|
|
137 |
"\" ( p <=> ( q <=> p ) ) \""} |
|
138 |
||
139 |
Let us check, what the solver returns when given a tautology: |
|
140 |
||
141 |
@{ML_response
|
|
142 |
"IffSolver.decide (translate @{term \"p = (q = p) = q\"})"
|
|
143 |
"true"} |
|
144 |
||
145 |
And here is what it returns for a formula which is not valid: |
|
146 |
||
147 |
@{ML_response
|
|
148 |
"IffSolver.decide (translate @{term \"p = (q = p)\"})"
|
|
149 |
"false"} |
|
150 |
*} |
|
| 61 | 151 |
|
| 79 | 152 |
text {*
|
| 94 | 153 |
Now, we combine these functions into an oracle. In general, an oracle may |
154 |
be given any input, but it has to return a certified proposition (a |
|
155 |
special term which is type-checked), out of which Isabelle's inference |
|
156 |
kernel ``magically'' makes a theorem. |
|
| 79 | 157 |
|
| 94 | 158 |
Here, we take the proposition to be show as input. Note that we have |
159 |
to first extract the term which is then passed to the translation and |
|
160 |
decision procedure. If the solver finds this term to be valid, we return |
|
161 |
the given proposition unchanged to be turned then into a theorem: |
|
| 61 | 162 |
*} |
163 |
||
| 94 | 164 |
oracle iff_oracle = {* fn ct =>
|
165 |
if IffSolver.decide (translate (HOLogic.dest_Trueprop (Thm.term_of ct))) |
|
166 |
then ct |
|
167 |
else error "Proof failed."*} |
|
| 61 | 168 |
|
| 94 | 169 |
text {*
|
170 |
Here is what we get when applying the oracle: |
|
171 |
||
172 |
@{ML_response_fake "iff_oracle @{cprop \"p = (p::bool)\"}" "p = p"}
|
|
173 |
||
174 |
(FIXME: is there a better way to present the theorem?) |
|
175 |
||
176 |
To be able to use our oracle for Isar proofs, we wrap it into a tactic: |
|
| 61 | 177 |
*} |
178 |
||
| 94 | 179 |
ML{*val iff_oracle_tac =
|
180 |
CSUBGOAL (fn (goal, i) => |
|
181 |
(case try iff_oracle goal of |
|
182 |
NONE => no_tac |
|
183 |
| SOME thm => rtac thm i))*} |
|
| 79 | 184 |
|
| 94 | 185 |
text {*
|
186 |
and create a new method solely based on this tactic: |
|
187 |
*} |
|
| 79 | 188 |
|
| 94 | 189 |
method_setup iff_oracle = {*
|
190 |
Method.no_args (Method.SIMPLE_METHOD' iff_oracle_tac) |
|
191 |
*} "Oracle-based decision procedure for chains of equivalences" |
|
192 |
||
193 |
text {*
|
|
194 |
(FIXME: what does @{ML "Method.SIMPLE_METHOD'"} do? ... what do you mean?)
|
|
| 79 | 195 |
|
| 94 | 196 |
Finally, we can test our oracle to prove some theorems: |
197 |
*} |
|
| 61 | 198 |
|
| 94 | 199 |
lemma "p = (p::bool)" |
200 |
by iff_oracle |
|
201 |
||
202 |
lemma "p = (q = p) = q" |
|
203 |
by iff_oracle |
|
| 79 | 204 |
|
205 |
||
206 |
text {*
|
|
| 94 | 207 |
(FIXME: say something about what the proof of the oracle is ... what do you mean?) |
| 83 | 208 |
*} |
209 |
||
210 |
||
| 61 | 211 |
end |