author | mengj |
Tue, 07 Mar 2006 04:01:25 +0100 | |
changeset 19199 | b338c218cc6e |
parent 19193 | 45c8db82893d |
child 19721 | 515f660c0ccb |
permissions | -rw-r--r-- |
17905
1574533861b1
Added files in order to use external ATPs as oracles and invoke these ATPs by calling Isabelle methods (currently "vampire" and "eprover").
mengj
parents:
diff
changeset
|
1 |
(* ID: $Id$ |
17907 | 2 |
Author: Jia Meng, NICTA |
17905
1574533861b1
Added files in order to use external ATPs as oracles and invoke these ATPs by calling Isabelle methods (currently "vampire" and "eprover").
mengj
parents:
diff
changeset
|
3 |
*) |
1574533861b1
Added files in order to use external ATPs as oracles and invoke these ATPs by calling Isabelle methods (currently "vampire" and "eprover").
mengj
parents:
diff
changeset
|
4 |
|
17958 | 5 |
header {* ATP setup (Vampire and E prover) *} |
17905
1574533861b1
Added files in order to use external ATPs as oracles and invoke these ATPs by calling Isabelle methods (currently "vampire" and "eprover").
mengj
parents:
diff
changeset
|
6 |
|
17958 | 7 |
theory ResAtpMethods |
8 |
imports Reconstruction |
|
9 |
uses |
|
10 |
"Tools/res_atp_provers.ML" |
|
11 |
("Tools/res_atp_methods.ML") |
|
17905
1574533861b1
Added files in order to use external ATPs as oracles and invoke these ATPs by calling Isabelle methods (currently "vampire" and "eprover").
mengj
parents:
diff
changeset
|
12 |
|
1574533861b1
Added files in order to use external ATPs as oracles and invoke these ATPs by calling Isabelle methods (currently "vampire" and "eprover").
mengj
parents:
diff
changeset
|
13 |
begin |
17939
3925ab7b8a18
Merged theory ResAtpOracle.thy into ResAtpMethods.thy
mengj
parents:
17907
diff
changeset
|
14 |
|
19193
45c8db82893d
When ATP methods call ATPs, all input clauses from one subgoal are written to one file.
mengj
parents:
18201
diff
changeset
|
15 |
oracle vampire_oracle ("string * int") = {* ResAtpProvers.vampire_o *} |
45c8db82893d
When ATP methods call ATPs, all input clauses from one subgoal are written to one file.
mengj
parents:
18201
diff
changeset
|
16 |
oracle eprover_oracle ("string * int") = {* ResAtpProvers.eprover_o *} |
17939
3925ab7b8a18
Merged theory ResAtpOracle.thy into ResAtpMethods.thy
mengj
parents:
17907
diff
changeset
|
17 |
|
17958 | 18 |
use "Tools/res_atp_methods.ML" |
19 |
setup ResAtpMethods.ResAtps_setup |
|
17939
3925ab7b8a18
Merged theory ResAtpOracle.thy into ResAtpMethods.thy
mengj
parents:
17907
diff
changeset
|
20 |
|
17907 | 21 |
end |