| author | wenzelm |
| Mon, 26 Apr 2010 21:50:28 +0200 | |
| changeset 36357 | 641a521bfc19 |
| parent 36064 | 48aec67c284f |
| child 36375 | 2482446a604c |
| permissions | -rw-r--r-- |
| 35827 | 1 |
(* Title: HOL/Sledgehammer.thy |
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
2 |
Author: Lawrence C Paulson |
|
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
3 |
Author: Jia Meng, NICTA |
| 35865 | 4 |
Author: Fabian Immler, TU Muenchen |
5 |
Author: Jasmin Blanchette, TU Muenchen |
|
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
6 |
*) |
|
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
7 |
|
| 35827 | 8 |
header {* Sledgehammer: Isabelle--ATP Linkup *}
|
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
9 |
|
| 35827 | 10 |
theory Sledgehammer |
| 33593 | 11 |
imports Plain Hilbert_Choice |
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
12 |
uses |
| 35865 | 13 |
"~~/src/Tools/Metis/metis.ML" |
| 35967 | 14 |
"Tools/Sledgehammer/sledgehammer_util.ML" |
| 35865 | 15 |
("Tools/Sledgehammer/sledgehammer_fol_clause.ML")
|
|
35825
a6aad5a70ed4
move Sledgehammer files in a directory of their own
blanchet
parents:
33593
diff
changeset
|
16 |
("Tools/Sledgehammer/sledgehammer_fact_preprocessor.ML")
|
|
a6aad5a70ed4
move Sledgehammer files in a directory of their own
blanchet
parents:
33593
diff
changeset
|
17 |
("Tools/Sledgehammer/sledgehammer_hol_clause.ML")
|
|
a6aad5a70ed4
move Sledgehammer files in a directory of their own
blanchet
parents:
33593
diff
changeset
|
18 |
("Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML")
|
|
a6aad5a70ed4
move Sledgehammer files in a directory of their own
blanchet
parents:
33593
diff
changeset
|
19 |
("Tools/Sledgehammer/sledgehammer_fact_filter.ML")
|
|
32327
0971cc0b6a57
src/HOL/Tools/ATP_Manager as separate component, with (almost) everything in one place;
wenzelm
parents:
31835
diff
changeset
|
20 |
("Tools/ATP_Manager/atp_manager.ML")
|
|
0971cc0b6a57
src/HOL/Tools/ATP_Manager as separate component, with (almost) everything in one place;
wenzelm
parents:
31835
diff
changeset
|
21 |
("Tools/ATP_Manager/atp_wrapper.ML")
|
|
0971cc0b6a57
src/HOL/Tools/ATP_Manager as separate component, with (almost) everything in one place;
wenzelm
parents:
31835
diff
changeset
|
22 |
("Tools/ATP_Manager/atp_minimal.ML")
|
|
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35865
diff
changeset
|
23 |
("Tools/Sledgehammer/sledgehammer_isar.ML")
|
| 35865 | 24 |
("Tools/Sledgehammer/meson_tactic.ML")
|
|
35825
a6aad5a70ed4
move Sledgehammer files in a directory of their own
blanchet
parents:
33593
diff
changeset
|
25 |
("Tools/Sledgehammer/metis_tactics.ML")
|
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
26 |
begin |
|
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
27 |
|
| 35865 | 28 |
definition COMBI :: "'a \<Rightarrow> 'a" |
29 |
where "COMBI P \<equiv> P" |
|
| 24819 | 30 |
|
| 35865 | 31 |
definition COMBK :: "'a \<Rightarrow> 'b \<Rightarrow> 'a" |
32 |
where "COMBK P Q \<equiv> P" |
|
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
33 |
|
| 35865 | 34 |
definition COMBB :: "('b => 'c) \<Rightarrow> ('a => 'b) \<Rightarrow> 'a \<Rightarrow> 'c"
|
35 |
where "COMBB P Q R \<equiv> P (Q R)" |
|
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
36 |
|
| 35865 | 37 |
definition COMBC :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c"
|
38 |
where "COMBC P Q R \<equiv> P R Q" |
|
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
39 |
|
| 35865 | 40 |
definition COMBS :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c"
|
41 |
where "COMBS P Q R \<equiv> P R (Q R)" |
|
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
42 |
|
| 35865 | 43 |
definition fequal :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
44 |
where "fequal X Y \<equiv> (X = Y)" |
|
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
45 |
|
| 35865 | 46 |
lemma fequal_imp_equal: "fequal X Y \<Longrightarrow> X = Y" |
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
47 |
by (simp add: fequal_def) |
|
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
48 |
|
| 35865 | 49 |
lemma equal_imp_fequal: "X = Y \<Longrightarrow> fequal X Y" |
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
50 |
by (simp add: fequal_def) |
|
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
51 |
|
|
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
52 |
text{*These two represent the equivalence between Boolean equality and iff.
|
|
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
53 |
They can't be converted to clauses automatically, as the iff would be |
|
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
54 |
expanded...*} |
|
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
55 |
|
| 35865 | 56 |
lemma iff_positive: "P \<or> Q \<or> P = Q" |
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
57 |
by blast |
|
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
58 |
|
| 35865 | 59 |
lemma iff_negative: "\<not> P \<or> \<not> Q \<or> P = Q" |
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
60 |
by blast |
|
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
61 |
|
| 24827 | 62 |
text{*Theorems for translation to combinators*}
|
63 |
||
| 35865 | 64 |
lemma abs_S: "\<lambda>x. (f x) (g x) \<equiv> COMBS f g" |
| 24827 | 65 |
apply (rule eq_reflection) |
66 |
apply (rule ext) |
|
67 |
apply (simp add: COMBS_def) |
|
68 |
done |
|
69 |
||
| 35865 | 70 |
lemma abs_I: "\<lambda>x. x \<equiv> COMBI" |
| 24827 | 71 |
apply (rule eq_reflection) |
72 |
apply (rule ext) |
|
73 |
apply (simp add: COMBI_def) |
|
74 |
done |
|
75 |
||
| 35865 | 76 |
lemma abs_K: "\<lambda>x. y \<equiv> COMBK y" |
| 24827 | 77 |
apply (rule eq_reflection) |
78 |
apply (rule ext) |
|
79 |
apply (simp add: COMBK_def) |
|
80 |
done |
|
81 |
||
| 35865 | 82 |
lemma abs_B: "\<lambda>x. a (g x) \<equiv> COMBB a g" |
| 24827 | 83 |
apply (rule eq_reflection) |
84 |
apply (rule ext) |
|
85 |
apply (simp add: COMBB_def) |
|
86 |
done |
|
87 |
||
| 35865 | 88 |
lemma abs_C: "\<lambda>x. (f x) b \<equiv> COMBC f b" |
| 24827 | 89 |
apply (rule eq_reflection) |
90 |
apply (rule ext) |
|
91 |
apply (simp add: COMBC_def) |
|
92 |
done |
|
93 |
||
|
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35865
diff
changeset
|
94 |
|
| 28585 | 95 |
subsection {* Setup of external ATPs *}
|
| 27368 | 96 |
|
| 35865 | 97 |
use "Tools/Sledgehammer/sledgehammer_fol_clause.ML" |
|
35825
a6aad5a70ed4
move Sledgehammer files in a directory of their own
blanchet
parents:
33593
diff
changeset
|
98 |
use "Tools/Sledgehammer/sledgehammer_fact_preprocessor.ML" |
| 35826 | 99 |
setup Sledgehammer_Fact_Preprocessor.setup |
|
35825
a6aad5a70ed4
move Sledgehammer files in a directory of their own
blanchet
parents:
33593
diff
changeset
|
100 |
use "Tools/Sledgehammer/sledgehammer_hol_clause.ML" |
|
a6aad5a70ed4
move Sledgehammer files in a directory of their own
blanchet
parents:
33593
diff
changeset
|
101 |
use "Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML" |
|
a6aad5a70ed4
move Sledgehammer files in a directory of their own
blanchet
parents:
33593
diff
changeset
|
102 |
use "Tools/Sledgehammer/sledgehammer_fact_filter.ML" |
| 35867 | 103 |
use "Tools/ATP_Manager/atp_manager.ML" |
|
35825
a6aad5a70ed4
move Sledgehammer files in a directory of their own
blanchet
parents:
33593
diff
changeset
|
104 |
use "Tools/ATP_Manager/atp_wrapper.ML" |
|
a6aad5a70ed4
move Sledgehammer files in a directory of their own
blanchet
parents:
33593
diff
changeset
|
105 |
setup ATP_Wrapper.setup |
|
32327
0971cc0b6a57
src/HOL/Tools/ATP_Manager as separate component, with (almost) everything in one place;
wenzelm
parents:
31835
diff
changeset
|
106 |
use "Tools/ATP_Manager/atp_minimal.ML" |
|
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35865
diff
changeset
|
107 |
use "Tools/Sledgehammer/sledgehammer_isar.ML" |
|
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35865
diff
changeset
|
108 |
|
| 35865 | 109 |
|
110 |
subsection {* The MESON prover *}
|
|
111 |
||
112 |
use "Tools/Sledgehammer/meson_tactic.ML" |
|
113 |
setup Meson_Tactic.setup |
|
| 28573 | 114 |
|
| 23444 | 115 |
|
116 |
subsection {* The Metis prover *}
|
|
117 |
||
|
35825
a6aad5a70ed4
move Sledgehammer files in a directory of their own
blanchet
parents:
33593
diff
changeset
|
118 |
use "Tools/Sledgehammer/metis_tactics.ML" |
| 35826 | 119 |
setup Metis_Tactics.setup |
| 23444 | 120 |
|
|
21254
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
wenzelm
parents:
diff
changeset
|
121 |
end |