author | blanchet |
Tue, 26 Oct 2010 14:06:21 +0200 | |
changeset 40185 | a6a34e0313bb |
parent 40181 | 3788b7adab36 |
child 41042 | 8275f52ac991 |
permissions | -rw-r--r-- |
35827 | 1 |
(* Title: HOL/Sledgehammer.thy |
38027 | 2 |
Author: Lawrence C. Paulson, Cambridge University Computer Laboratory |
38028 | 3 |
Author: Jia Meng, Cambridge University Computer Laboratory and NICTA |
35865 | 4 |
Author: Jasmin Blanchette, TU Muenchen |
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
5 |
*) |
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
6 |
|
35827 | 7 |
header {* Sledgehammer: Isabelle--ATP Linkup *} |
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
8 |
|
35827 | 9 |
theory Sledgehammer |
40181 | 10 |
imports ATP SMT |
39951
ff60a6e4edfe
factor out "ATP" from "Sledgehammer" (cf. "SAT" vs. "Refute", etc.) -- the theories now reflect the directory structure
blanchet
parents:
39947
diff
changeset
|
11 |
uses "Tools/Sledgehammer/sledgehammer_util.ML" |
39947 | 12 |
"Tools/Sledgehammer/sledgehammer_filter.ML" |
40067 | 13 |
"Tools/Sledgehammer/sledgehammer_atp_translate.ML" |
14 |
"Tools/Sledgehammer/sledgehammer_atp_reconstruct.ML" |
|
39947 | 15 |
"Tools/Sledgehammer/sledgehammer.ML" |
16 |
"Tools/Sledgehammer/sledgehammer_minimize.ML" |
|
17 |
"Tools/Sledgehammer/sledgehammer_isar.ML" |
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
18 |
begin |
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
19 |
|
39947 | 20 |
setup {* |
39951
ff60a6e4edfe
factor out "ATP" from "Sledgehammer" (cf. "SAT" vs. "Refute", etc.) -- the theories now reflect the directory structure
blanchet
parents:
39947
diff
changeset
|
21 |
Sledgehammer.setup |
39947 | 22 |
#> Sledgehammer_Isar.setup |
23 |
*} |
|
23444 | 24 |
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
25 |
end |