61
|
1 |
theory ExternalSolver
|
346
|
2 |
imports "../Appendix"
|
61
|
3 |
begin
|
|
4 |
|
94
|
5 |
|
565
|
6 |
section \<open>Executing an External Application (TBD) \label{rec:external}\<close>
|
61
|
7 |
|
565
|
8 |
text \<open>
|
61
|
9 |
{\bf Problem:}
|
94
|
10 |
You want to use an external application.
|
|
11 |
\smallskip
|
|
12 |
|
472
|
13 |
{\bf Solution:} The function @{ML bash_output in Isabelle_System} might be the right thing for
|
94
|
14 |
you.
|
61
|
15 |
\smallskip
|
|
16 |
|
94
|
17 |
This function executes an external command as if printed in a shell. It
|
|
18 |
returns the output of the program and its return value.
|
|
19 |
|
|
20 |
For example, consider running an ordinary shell commands:
|
|
21 |
|
567
|
22 |
@{ML_matchresult [display,gray]
|
569
|
23 |
\<open>Isabelle_System.bash_output "echo Hello world!"\<close> \<open>("Hello world!\n", 0)\<close>}
|
94
|
24 |
|
|
25 |
Note that it works also fine with timeouts (see Recipe~\ref{rec:timeout}
|
|
26 |
on Page~\pageref{rec:timeout}), i.e. external applications are killed
|
|
27 |
properly. For example, the following expression takes only approximately
|
|
28 |
one second:
|
|
29 |
|
572
|
30 |
@{ML_response [display,gray]
|
569
|
31 |
\<open>Timeout.apply (Time.fromSeconds 1) Isabelle_System.bash_output "sleep 30"
|
|
32 |
handle TIMEOUT => ("timeout", ~1)\<close> \<open>("timeout", ~1)\<close>}
|
94
|
33 |
|
472
|
34 |
The function @{ML bash_output in Isabelle_System} can also be used for more reasonable
|
94
|
35 |
applications, e.g. coupling external solvers with Isabelle. In that case,
|
|
36 |
one has to make sure that Isabelle can find the particular executable.
|
|
37 |
One way to ensure this is by adding a Bash-like variable binding into
|
|
38 |
one of Isabelle's settings file (prefer the user settings file usually to
|
565
|
39 |
be found at \<open>$HOME/.isabelle/etc/settings\<close>).
|
94
|
40 |
|
565
|
41 |
For example, assume you want to use the application \<open>foo\<close> which
|
|
42 |
is here supposed to be located at \<open>/usr/local/bin/\<close>.
|
94
|
43 |
The following line has to be added to one of Isabelle's settings file:
|
|
44 |
|
565
|
45 |
\<open>FOO=/usr/local/bin/foo\<close>
|
94
|
46 |
|
|
47 |
In Isabelle, this application may now be executed by
|
|
48 |
|
572
|
49 |
@{ML_response [display,gray] \<open>Isabelle_System.bash_output "$FOO"\<close> \<open>\<dots>\<close>}
|
565
|
50 |
\<close>
|
94
|
51 |
|
|
52 |
|
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
53 |
end
|