author | wenzelm |
Thu, 29 Jan 2015 16:35:29 +0100 | |
changeset 59471 | ca459352d8c5 |
parent 58889 | 5b7a9633cfa8 |
child 60758 | d8d85a8172b5 |
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 |
|
58889 | 7 |
section {* Sledgehammer: Isabelle--ATP Linkup *} |
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
8 |
|
35827 | 9 |
theory Sledgehammer |
58061 | 10 |
imports Presburger SMT |
46950
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents:
43085
diff
changeset
|
11 |
keywords "sledgehammer" :: diag and "sledgehammer_params" :: thy_decl |
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
12 |
begin |
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
13 |
|
54838
16511f84913c
reconstruct SPASS-Pirate steps of the form 'x ~= C x' (or more complicated)
blanchet
parents:
52641
diff
changeset
|
14 |
lemma size_ne_size_imp_ne: "size x \<noteq> size y \<Longrightarrow> x \<noteq> y" |
58481 | 15 |
by (erule contrapos_nn) (rule arg_cong) |
54838
16511f84913c
reconstruct SPASS-Pirate steps of the form 'x ~= C x' (or more complicated)
blanchet
parents:
52641
diff
changeset
|
16 |
|
59471
ca459352d8c5
more explicit indication of Async_Manager_Legacy as Proof General legacy;
wenzelm
parents:
58889
diff
changeset
|
17 |
ML_file "Tools/Sledgehammer/async_manager_legacy.ML" |
48891 | 18 |
ML_file "Tools/Sledgehammer/sledgehammer_util.ML" |
19 |
ML_file "Tools/Sledgehammer/sledgehammer_fact.ML" |
|
55287 | 20 |
ML_file "Tools/Sledgehammer/sledgehammer_proof_methods.ML" |
55211 | 21 |
ML_file "Tools/Sledgehammer/sledgehammer_isar_annotate.ML" |
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
22 |
ML_file "Tools/Sledgehammer/sledgehammer_isar_proof.ML" |
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
23 |
ML_file "Tools/Sledgehammer/sledgehammer_isar_preplay.ML" |
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
24 |
ML_file "Tools/Sledgehammer/sledgehammer_isar_compress.ML" |
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
25 |
ML_file "Tools/Sledgehammer/sledgehammer_isar_minimize.ML" |
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
26 |
ML_file "Tools/Sledgehammer/sledgehammer_isar.ML" |
55201 | 27 |
ML_file "Tools/Sledgehammer/sledgehammer_prover.ML" |
55205 | 28 |
ML_file "Tools/Sledgehammer/sledgehammer_prover_atp.ML" |
58061 | 29 |
ML_file "Tools/Sledgehammer/sledgehammer_prover_smt.ML" |
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
30 |
ML_file "Tools/Sledgehammer/sledgehammer_prover_minimize.ML" |
48891 | 31 |
ML_file "Tools/Sledgehammer/sledgehammer_mepo.ML" |
32 |
ML_file "Tools/Sledgehammer/sledgehammer_mash.ML" |
|
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
33 |
ML_file "Tools/Sledgehammer/sledgehammer.ML" |
55198 | 34 |
ML_file "Tools/Sledgehammer/sledgehammer_commands.ML" |
48891 | 35 |
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
36 |
end |