ProgTutorial/Recipes/ExternalSolver.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 02 Jun 2010 10:08:16 +0200
changeset 437 e2b351efd6f3
parent 418 1d1e4cda8c54
child 472 1bbe4268664d
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     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
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     3
begin
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     4
94
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
     5
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
     6
section {* Executing an External Application (TBD) \label{rec:external}*}
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     7
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     8
text {*
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     9
  {\bf Problem:}
94
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    10
  You want to use an external application.
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    11
  \smallskip
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    12
418
1d1e4cda8c54 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 349
diff changeset
    13
  {\bf Solution:} The function @{ML bash_output} might be the right thing for
94
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    14
  you.
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    15
  \smallskip
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    16
94
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    17
  This function executes an external command as if printed in a shell. It
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    18
  returns the output of the program and its return value.
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    19
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    20
  For example, consider running an ordinary shell commands:
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    21
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    22
  @{ML_response [display,gray] 
418
1d1e4cda8c54 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 349
diff changeset
    23
    "bash_output \"echo Hello world!\"" "(\"Hello world!\\n\", 0)"}
94
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    24
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    25
  Note that it works also fine with timeouts (see Recipe~\ref{rec:timeout}
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    26
  on Page~\pageref{rec:timeout}), i.e. external applications are killed
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    27
  properly. For example, the following expression takes only approximately
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    28
  one second:
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    29
349
9e374cd891e1 updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
    30
  @{ML_response_fake [display,gray] 
418
1d1e4cda8c54 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 349
diff changeset
    31
    "TimeLimit.timeLimit (Time.fromSeconds 1) bash_output \"sleep 30\"
94
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    32
     handle TimeLimit.TimeOut => (\"timeout\", ~1)" "(\"timeout\", ~1)"}
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    33
418
1d1e4cda8c54 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 349
diff changeset
    34
  The function @{ML bash_output} can also be used for more reasonable
94
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    35
  applications, e.g. coupling external solvers with Isabelle. In that case,
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    36
  one has to make sure that Isabelle can find the particular executable.
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    37
  One way to ensure this is by adding a Bash-like variable binding into
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    38
  one of Isabelle's settings file (prefer the user settings file usually to
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    39
  be found at @{text "$HOME/.isabelle/etc/settings"}).
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    40
  
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    41
  For example, assume you want to use the application @{text foo} which
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    42
  is here supposed to be located at @{text "/usr/local/bin/"}.
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    43
  The following line has to be added to one of Isabelle's settings file:
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    44
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    45
  @{text "FOO=/usr/local/bin/foo"}
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    46
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    47
  In Isabelle, this application may now be executed by
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    48
418
1d1e4cda8c54 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 349
diff changeset
    49
  @{ML_response_fake [display,gray] "bash_output \"$FOO\"" "\<dots>"}
94
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    50
*}
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    51
531e453c9d67 rewrote recipes describing external solvers
boehmes
parents: 83
diff changeset
    52
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    53
end