author | huffman |
Thu, 27 Oct 2011 07:46:57 +0200 | |
changeset 45270 | d5b5c9259afd |
parent 43085 | 0a2f5b86bdd7 |
child 46950 | d0181abdbdac |
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 |
41042
8275f52ac991
load "try" after "Metis" and move "Async_Manager" back to Sledgehammer
blanchet
parents:
40181
diff
changeset
|
11 |
uses "Tools/Sledgehammer/async_manager.ML" |
8275f52ac991
load "try" after "Metis" and move "Async_Manager" back to Sledgehammer
blanchet
parents:
40181
diff
changeset
|
12 |
"Tools/Sledgehammer/sledgehammer_util.ML" |
39947 | 13 |
"Tools/Sledgehammer/sledgehammer_filter.ML" |
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41042
diff
changeset
|
14 |
"Tools/Sledgehammer/sledgehammer_provers.ML" |
39947 | 15 |
"Tools/Sledgehammer/sledgehammer_minimize.ML" |
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41042
diff
changeset
|
16 |
"Tools/Sledgehammer/sledgehammer_run.ML" |
39947 | 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 |
|
42616
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents:
41094
diff
changeset
|
20 |
setup {* Sledgehammer_Isar.setup *} |
23444 | 21 |
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
22 |
end |