author | Norbert Schirmer <norbert.schirmer@web.de> |
Thu, 16 May 2019 19:56:12 +0200 | |
changeset 567 | f7c97e64cc2a |
parent 565 | cecd7a941885 |
child 569 | f875a25aa72d |
permissions | -rw-r--r-- |
61 | 1 |
theory ExternalSolver |
346
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents:
211
diff
changeset
|
2 |
imports "../Appendix" |
61 | 3 |
begin |
4 |
||
94 | 5 |
|
565
cecd7a941885
isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents:
562
diff
changeset
|
6 |
section \<open>Executing an External Application (TBD) \label{rec:external}\<close> |
61 | 7 |
|
565
cecd7a941885
isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents:
562
diff
changeset
|
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
f7c97e64cc2a
tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents:
565
diff
changeset
|
22 |
@{ML_matchresult [display,gray] |
472 | 23 |
"Isabelle_System.bash_output \"echo Hello world!\"" "(\"Hello world!\\n\", 0)"} |
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 |
||
567
f7c97e64cc2a
tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents:
565
diff
changeset
|
30 |
@{ML_matchresult_fake [display,gray] |
562
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
517
diff
changeset
|
31 |
"Timeout.apply (Time.fromSeconds 1) Isabelle_System.bash_output \"sleep 30\" |
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
517
diff
changeset
|
32 |
handle TIMEOUT => (\"timeout\", ~1)" "(\"timeout\", ~1)"} |
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
cecd7a941885
isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents:
562
diff
changeset
|
39 |
be found at \<open>$HOME/.isabelle/etc/settings\<close>). |
94 | 40 |
|
565
cecd7a941885
isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents:
562
diff
changeset
|
41 |
For example, assume you want to use the application \<open>foo\<close> which |
cecd7a941885
isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents:
562
diff
changeset
|
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
cecd7a941885
isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents:
562
diff
changeset
|
45 |
\<open>FOO=/usr/local/bin/foo\<close> |
94 | 46 |
|
47 |
In Isabelle, this application may now be executed by |
|
48 |
||
567
f7c97e64cc2a
tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents:
565
diff
changeset
|
49 |
@{ML_matchresult_fake [display,gray] "Isabelle_System.bash_output \"$FOO\"" "\<dots>"} |
565
cecd7a941885
isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents:
562
diff
changeset
|
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>
parents:
472
diff
changeset
|
53 |
end |