| author | haftmann | 
| Tue, 08 Jun 2010 16:37:22 +0200 | |
| changeset 37388 | 793618618f78 | 
| parent 36569 | 3a29eb7606c3 | 
| child 37399 | 34f080a12063 | 
| 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")
 | 
| 
36377
 
b3dce4c715d0
now rename the file "atp_wrapper.ML" to "atp_systems.ML" + fix typo in "SystemOnTPTP" script
 
blanchet 
parents: 
36376 
diff
changeset
 | 
21  | 
  ("Tools/ATP_Manager/atp_systems.ML")
 | 
| 
36375
 
2482446a604c
move the minimizer to the Sledgehammer directory
 
blanchet 
parents: 
36064 
diff
changeset
 | 
22  | 
  ("Tools/Sledgehammer/sledgehammer_fact_minimizer.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  | 
|
| 
36569
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
28  | 
definition COMBI :: "'a \<Rightarrow> 'a" where  | 
| 
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
29  | 
[no_atp]: "COMBI P \<equiv> P"  | 
| 24819 | 30  | 
|
| 
36569
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
31  | 
definition COMBK :: "'a \<Rightarrow> 'b \<Rightarrow> 'a" where  | 
| 
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
32  | 
[no_atp]: "COMBK P Q \<equiv> P"  | 
| 
21254
 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 
wenzelm 
parents:  
diff
changeset
 | 
33  | 
|
| 
36569
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
34  | 
definition COMBB :: "('b => 'c) \<Rightarrow> ('a => 'b) \<Rightarrow> 'a \<Rightarrow> 'c" where [no_atp]:
 | 
| 
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
35  | 
"COMBB P Q R \<equiv> P (Q R)"  | 
| 
21254
 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 
wenzelm 
parents:  
diff
changeset
 | 
36  | 
|
| 
36569
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
37  | 
definition COMBC :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c" where
 | 
| 
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
38  | 
[no_atp]: "COMBC P Q R \<equiv> P R Q"  | 
| 
21254
 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 
wenzelm 
parents:  
diff
changeset
 | 
39  | 
|
| 
36569
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
40  | 
definition COMBS :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" where
 | 
| 
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
41  | 
[no_atp]: "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  | 
|
| 
36569
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
43  | 
definition fequal :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where [no_atp]:  | 
| 
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
44  | 
"fequal X Y \<equiv> (X = Y)"  | 
| 
21254
 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
|
| 
36569
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
46  | 
lemma fequal_imp_equal [no_atp]: "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  | 
|
| 
36569
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
49  | 
lemma equal_imp_fequal [no_atp]: "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  | 
||
| 
36569
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
64  | 
lemma abs_S [no_atp]: "\<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  | 
||
| 
36569
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
70  | 
lemma abs_I [no_atp]: "\<lambda>x. x \<equiv> COMBI"  | 
| 24827 | 71  | 
apply (rule eq_reflection)  | 
72  | 
apply (rule ext)  | 
|
73  | 
apply (simp add: COMBI_def)  | 
|
74  | 
done  | 
|
75  | 
||
| 
36569
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
76  | 
lemma abs_K [no_atp]: "\<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  | 
||
| 
36569
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
82  | 
lemma abs_B [no_atp]: "\<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  | 
||
| 
36569
 
3a29eb7606c3
added "no_atp" for theorems that are automatically used or included by Sledgehammer when appropriate (about combinators and fequal)
 
blanchet 
parents: 
36394 
diff
changeset
 | 
88  | 
lemma abs_C [no_atp]: "\<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"  | 
| 
36377
 
b3dce4c715d0
now rename the file "atp_wrapper.ML" to "atp_systems.ML" + fix typo in "SystemOnTPTP" script
 
blanchet 
parents: 
36376 
diff
changeset
 | 
104  | 
use "Tools/ATP_Manager/atp_systems.ML"  | 
| 36376 | 105  | 
setup ATP_Systems.setup  | 
| 
36375
 
2482446a604c
move the minimizer to the Sledgehammer directory
 
blanchet 
parents: 
36064 
diff
changeset
 | 
106  | 
use "Tools/Sledgehammer/sledgehammer_fact_minimizer.ML"  | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents: 
35865 
diff
changeset
 | 
107  | 
use "Tools/Sledgehammer/sledgehammer_isar.ML"  | 
| 
36394
 
1a48d18449d8
move "neg_clausify" method and "clausify" attribute to "sledgehammer_isar.ML"
 
blanchet 
parents: 
36377 
diff
changeset
 | 
108  | 
setup Sledgehammer_Isar.setup  | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents: 
35865 
diff
changeset
 | 
109  | 
|
| 35865 | 110  | 
|
111  | 
subsection {* The MESON prover *}
 | 
|
112  | 
||
113  | 
use "Tools/Sledgehammer/meson_tactic.ML"  | 
|
114  | 
setup Meson_Tactic.setup  | 
|
| 28573 | 115  | 
|
| 23444 | 116  | 
|
117  | 
subsection {* The Metis prover *}
 | 
|
118  | 
||
| 
35825
 
a6aad5a70ed4
move Sledgehammer files in a directory of their own
 
blanchet 
parents: 
33593 
diff
changeset
 | 
119  | 
use "Tools/Sledgehammer/metis_tactics.ML"  | 
| 35826 | 120  | 
setup Metis_Tactics.setup  | 
| 23444 | 121  | 
|
| 
21254
 
d53f76357f41
incorporated former theories Reconstruction and ResAtpMethods into ATP_Linkup;
 
wenzelm 
parents:  
diff
changeset
 | 
122  | 
end  |