| author | wenzelm | 
| Thu, 18 Jul 2024 17:02:39 +0200 | |
| changeset 80587 | 12de235f8b92 | 
| parent 79734 | 0fa4bebbdd75 | 
| child 80910 | 406a85a25189 | 
| permissions | -rw-r--r-- | 
| 39452 | 1  | 
(* Title: HOL/Tools/ATP/atp_proof.ML  | 
2  | 
Author: Lawrence C. Paulson, Cambridge University Computer Laboratory  | 
|
3  | 
Author: Claire Quigley, Cambridge University Computer Laboratory  | 
|
4  | 
Author: Jasmin Blanchette, TU Muenchen  | 
|
| 
57697
 
44341963ade3
correctly translate THF functions from terms to types
 
blanchet 
parents: 
57656 
diff
changeset
 | 
5  | 
Author: Mathias Fleury, ENS Rennes  | 
| 39452 | 6  | 
|
| 
42876
 
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
 
blanchet 
parents: 
42848 
diff
changeset
 | 
7  | 
Abstract representation of ATP proofs and TSTP/SPASS syntax.  | 
| 39452 | 8  | 
*)  | 
9  | 
||
10  | 
signature ATP_PROOF =  | 
|
11  | 
sig  | 
|
| 54811 | 12  | 
type 'a atp_type = 'a ATP_Problem.atp_type  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
13  | 
  type ('a, 'b) atp_term = ('a, 'b) ATP_Problem.atp_term
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
14  | 
type atp_formula_role = ATP_Problem.atp_formula_role  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
15  | 
  type ('a, 'b, 'c, 'd) atp_formula = ('a, 'b, 'c, 'd) ATP_Problem.atp_formula
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
16  | 
type 'a atp_problem = 'a ATP_Problem.atp_problem  | 
| 39452 | 17  | 
|
| 
42965
 
1403595ec38c
slightly gracefuller handling of LEO-II and Satallax output
 
blanchet 
parents: 
42962 
diff
changeset
 | 
18  | 
exception UNRECOGNIZED_ATP_PROOF of unit  | 
| 
 
1403595ec38c
slightly gracefuller handling of LEO-II and Satallax output
 
blanchet 
parents: 
42962 
diff
changeset
 | 
19  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
20  | 
datatype atp_failure =  | 
| 
58654
 
3e1cad27fc2f
special treatment of extensionality in minimizer
 
blanchet 
parents: 
58600 
diff
changeset
 | 
21  | 
MaybeUnprovable |  | 
| 
42587
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
22  | 
Unprovable |  | 
| 
43050
 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 
blanchet 
parents: 
43029 
diff
changeset
 | 
23  | 
GaveUp |  | 
| 
42587
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
24  | 
ProofMissing |  | 
| 
42882
 
391e41ac038b
make sure the Vampire incomplete proof detection code kicks in
 
blanchet 
parents: 
42876 
diff
changeset
 | 
25  | 
ProofIncomplete |  | 
| 
57266
 
6a3b5085fb8f
fixed parsing of one-argument 'file()' in TSTP files
 
blanchet 
parents: 
57261 
diff
changeset
 | 
26  | 
ProofUnparsable |  | 
| 
77430
 
51dac6fcdd0e
reverted 0506c3273814 -- the message is still useful
 
blanchet 
parents: 
77427 
diff
changeset
 | 
27  | 
UnsoundProof of bool * string list |  | 
| 
42587
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
28  | 
TimedOut |  | 
| 
42953
 
26111aafab12
detect inappropriate problems and crashes better in Waldmeister
 
blanchet 
parents: 
42943 
diff
changeset
 | 
29  | 
Inappropriate |  | 
| 
42587
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
30  | 
OutOfResources |  | 
| 
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
31  | 
MalformedInput |  | 
| 
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
32  | 
MalformedOutput |  | 
| 
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
33  | 
Interrupted |  | 
| 
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
34  | 
Crashed |  | 
| 
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
35  | 
InternalError |  | 
| 
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
36  | 
UnknownError of string  | 
| 
39491
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
37  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
38  | 
type atp_step_name = string * string list  | 
| 
53587
 
3fb81ab13ea3
generalized data structure, for extension with SMT solver proofs
 
blanchet 
parents: 
53586 
diff
changeset
 | 
39  | 
  type ('a, 'b) atp_step =
 | 
| 
 
3fb81ab13ea3
generalized data structure, for extension with SMT solver proofs
 
blanchet 
parents: 
53586 
diff
changeset
 | 
40  | 
atp_step_name * atp_formula_role * 'a * 'b * atp_step_name list  | 
| 39452 | 41  | 
|
| 54811 | 42  | 
  type 'a atp_proof = (('a, 'a, ('a, 'a atp_type) atp_term, 'a) atp_formula, string) atp_step list
 | 
| 39452 | 43  | 
|
| 57154 | 44  | 
val agsyholN : string  | 
45  | 
val alt_ergoN : string  | 
|
| 75029 | 46  | 
val cvc4N : string  | 
| 57154 | 47  | 
val eN : string  | 
48  | 
val iproverN : string  | 
|
49  | 
val leo2N : string  | 
|
| 
67021
 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 
blanchet 
parents: 
66545 
diff
changeset
 | 
50  | 
val leo3N : string  | 
| 57154 | 51  | 
val satallaxN : string  | 
52  | 
val spassN : string  | 
|
53  | 
val vampireN : string  | 
|
| 75029 | 54  | 
val veritN : string  | 
| 57154 | 55  | 
val waldmeisterN : string  | 
| 75029 | 56  | 
val z3N : string  | 
| 57154 | 57  | 
val zipperpositionN : string  | 
58  | 
val remote_prefix : string  | 
|
| 74117 | 59  | 
val dummy_fofN : string  | 
| 72588 | 60  | 
val dummy_tfxN : string  | 
| 74109 | 61  | 
val dummy_thfN : string  | 
| 57154 | 62  | 
|
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
63  | 
val agsyhol_core_rule : string  | 
| 
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
64  | 
val spass_input_rule : string  | 
| 
55192
 
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
 
blanchet 
parents: 
54836 
diff
changeset
 | 
65  | 
val spass_pre_skolemize_rule : string  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
66  | 
val spass_skolemize_rule : string  | 
| 75122 | 67  | 
val zipperposition_define_rule : string  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
68  | 
|
| 
41259
 
13972ced98d9
more precise error messages in "verbose" (or "debug") mode, following this morning's permission debacle
 
blanchet 
parents: 
41222 
diff
changeset
 | 
69  | 
val short_output : bool -> string -> string  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
70  | 
val string_of_atp_failure : atp_failure -> string  | 
| 54811 | 71  | 
val extract_known_atp_failure : (atp_failure * string) list -> string -> atp_failure option  | 
| 
39491
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
72  | 
val extract_tstplike_proof_and_outcome :  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
73  | 
bool -> (string * string) list -> (atp_failure * string) list -> string  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
74  | 
-> string * atp_failure option  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
75  | 
val is_same_atp_step : atp_step_name -> atp_step_name -> bool  | 
| 42961 | 76  | 
val scan_general_id : string list -> string * string list  | 
| 72398 | 77  | 
val parse_fol_formula : string list ->  | 
| 54811 | 78  | 
(string, string atp_type, (string, string atp_type) atp_term, string) atp_formula * string list  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
79  | 
val clean_up_atp_proof_dependencies : string atp_proof -> string atp_proof  | 
| 54811 | 80  | 
val map_term_names_in_atp_proof : (string -> string) -> string atp_proof -> string atp_proof  | 
81  | 
val nasty_atp_proof : string Symtab.table -> string atp_proof -> string atp_proof  | 
|
| 
57707
 
0242e9578828
imported patch satallax_proof_support_Sledgehammer
 
fleury 
parents: 
57697 
diff
changeset
 | 
82  | 
|
| 
72399
 
f8900a5ad4a7
drop obsolete ad hoc support for Satallax isar proof reconstruction
 
desharna 
parents: 
72398 
diff
changeset
 | 
83  | 
val skip_term : string list -> string * string list  | 
| 
 
f8900a5ad4a7
drop obsolete ad hoc support for Satallax isar proof reconstruction
 
desharna 
parents: 
72398 
diff
changeset
 | 
84  | 
val parse_hol_formula : string list ->  | 
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57709 
diff
changeset
 | 
85  | 
    ('a, 'b, (string, string ATP_Problem.atp_type) ATP_Problem.atp_term, 'c) ATP_Problem.atp_formula *
 | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57709 
diff
changeset
 | 
86  | 
string list  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57709 
diff
changeset
 | 
87  | 
val dummy_atype : string ATP_Problem.atp_type  | 
| 
72399
 
f8900a5ad4a7
drop obsolete ad hoc support for Satallax isar proof reconstruction
 
desharna 
parents: 
72398 
diff
changeset
 | 
88  | 
val role_of_tptp_string : string -> ATP_Problem.atp_formula_role  | 
| 75064 | 89  | 
  val parse_line : bool -> string -> ('a * string ATP_Problem.atp_problem_line list) list ->
 | 
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57709 
diff
changeset
 | 
90  | 
string list -> ((string * string list) * ATP_Problem.atp_formula_role *  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57709 
diff
changeset
 | 
91  | 
(string, 'b, (string, string ATP_Problem.atp_type) ATP_Problem.atp_term,  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57709 
diff
changeset
 | 
92  | 
'c) ATP_Problem.atp_formula  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57709 
diff
changeset
 | 
93  | 
* string * (string * 'd list) list) list * string list  | 
| 57716 | 94  | 
  val core_inference : 'a -> 'b -> ('b * 'b list) * ATP_Problem.atp_formula_role *
 | 
95  | 
    ('c, 'd, (string, 'e) ATP_Problem.atp_term, 'f) ATP_Problem.atp_formula * 'a * 'g list
 | 
|
| 70586 | 96  | 
val vampire_step_name_ord : (string * 'a) ord  | 
| 
57707
 
0242e9578828
imported patch satallax_proof_support_Sledgehammer
 
fleury 
parents: 
57697 
diff
changeset
 | 
97  | 
val core_of_agsyhol_proof : string -> string list option  | 
| 72355 | 98  | 
  val string_of_atp_step : ('a -> string) -> ('b -> string) -> ('a, 'b) atp_step -> string
 | 
| 
72399
 
f8900a5ad4a7
drop obsolete ad hoc support for Satallax isar proof reconstruction
 
desharna 
parents: 
72398 
diff
changeset
 | 
99  | 
|
| 75064 | 100  | 
val atp_proof_of_tstplike_proof : bool -> string -> string atp_problem -> string ->  | 
101  | 
string atp_proof  | 
|
| 
77418
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
102  | 
val atp_abduce_candidates_of_output : string -> string atp_problem -> string ->  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
103  | 
(string, string, (string, string atp_type) atp_term, string) atp_formula list  | 
| 39452 | 104  | 
end;  | 
105  | 
||
106  | 
structure ATP_Proof : ATP_PROOF =  | 
|
107  | 
struct  | 
|
108  | 
||
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43050 
diff
changeset
 | 
109  | 
open ATP_Util  | 
| 
39491
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
110  | 
open ATP_Problem  | 
| 
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
111  | 
|
| 57154 | 112  | 
val agsyholN = "agsyhol"  | 
113  | 
val alt_ergoN = "alt_ergo"  | 
|
| 75029 | 114  | 
val cvc4N = "cvc4"  | 
| 57154 | 115  | 
val eN = "e"  | 
116  | 
val iproverN = "iprover"  | 
|
117  | 
val leo2N = "leo2"  | 
|
| 
67021
 
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
 
blanchet 
parents: 
66545 
diff
changeset
 | 
118  | 
val leo3N = "leo3"  | 
| 57154 | 119  | 
val satallaxN = "satallax"  | 
120  | 
val spassN = "spass"  | 
|
121  | 
val vampireN = "vampire"  | 
|
| 75029 | 122  | 
val veritN = "verit"  | 
| 57154 | 123  | 
val waldmeisterN = "waldmeister"  | 
| 75029 | 124  | 
val z3N = "z3"  | 
| 57154 | 125  | 
val zipperpositionN = "zipperposition"  | 
126  | 
val remote_prefix = "remote_"  | 
|
| 75029 | 127  | 
|
| 74117 | 128  | 
val dummy_fofN = "dummy_fof"  | 
| 72588 | 129  | 
val dummy_tfxN = "dummy_tfx"  | 
| 74109 | 130  | 
val dummy_thfN = "dummy_thf"  | 
| 57154 | 131  | 
|
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
132  | 
val agsyhol_core_rule = "__agsyhol_core" (* arbitrary *)  | 
| 
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
133  | 
val spass_input_rule = "Inp"  | 
| 
55192
 
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
 
blanchet 
parents: 
54836 
diff
changeset
 | 
134  | 
val spass_pre_skolemize_rule = "__Sko0" (* arbitrary *)  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
135  | 
val spass_skolemize_rule = "__Sko" (* arbitrary *)  | 
| 75122 | 136  | 
val zipperposition_define_rule = "define" (* arbitrary *)  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
137  | 
|
| 
42965
 
1403595ec38c
slightly gracefuller handling of LEO-II and Satallax output
 
blanchet 
parents: 
42962 
diff
changeset
 | 
138  | 
exception UNRECOGNIZED_ATP_PROOF of unit  | 
| 
 
1403595ec38c
slightly gracefuller handling of LEO-II and Satallax output
 
blanchet 
parents: 
42962 
diff
changeset
 | 
139  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
140  | 
datatype atp_failure =  | 
| 
58654
 
3e1cad27fc2f
special treatment of extensionality in minimizer
 
blanchet 
parents: 
58600 
diff
changeset
 | 
141  | 
MaybeUnprovable |  | 
| 
42587
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
142  | 
Unprovable |  | 
| 
43050
 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 
blanchet 
parents: 
43029 
diff
changeset
 | 
143  | 
GaveUp |  | 
| 
42587
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
144  | 
ProofMissing |  | 
| 
42882
 
391e41ac038b
make sure the Vampire incomplete proof detection code kicks in
 
blanchet 
parents: 
42876 
diff
changeset
 | 
145  | 
ProofIncomplete |  | 
| 
57266
 
6a3b5085fb8f
fixed parsing of one-argument 'file()' in TSTP files
 
blanchet 
parents: 
57261 
diff
changeset
 | 
146  | 
ProofUnparsable |  | 
| 
77430
 
51dac6fcdd0e
reverted 0506c3273814 -- the message is still useful
 
blanchet 
parents: 
77427 
diff
changeset
 | 
147  | 
UnsoundProof of bool * string list |  | 
| 
42587
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
148  | 
TimedOut |  | 
| 
42953
 
26111aafab12
detect inappropriate problems and crashes better in Waldmeister
 
blanchet 
parents: 
42943 
diff
changeset
 | 
149  | 
Inappropriate |  | 
| 
42587
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
150  | 
OutOfResources |  | 
| 
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
151  | 
MalformedInput |  | 
| 
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
152  | 
MalformedOutput |  | 
| 
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
153  | 
Interrupted |  | 
| 
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
154  | 
Crashed |  | 
| 
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
155  | 
InternalError |  | 
| 
 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 
blanchet 
parents: 
42550 
diff
changeset
 | 
156  | 
UnknownError of string  | 
| 
39491
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
157  | 
|
| 
41259
 
13972ced98d9
more precise error messages in "verbose" (or "debug") mode, following this morning's permission debacle
 
blanchet 
parents: 
41222 
diff
changeset
 | 
158  | 
fun short_output verbose output =  | 
| 
42060
 
889d767ce5f4
make Minimizer honor "verbose" and "debug" options better
 
blanchet 
parents: 
41944 
diff
changeset
 | 
159  | 
if verbose then  | 
| 
 
889d767ce5f4
make Minimizer honor "verbose" and "debug" options better
 
blanchet 
parents: 
41944 
diff
changeset
 | 
160  | 
if output = "" then "No details available" else elide_string 1000 output  | 
| 
 
889d767ce5f4
make Minimizer honor "verbose" and "debug" options better
 
blanchet 
parents: 
41944 
diff
changeset
 | 
161  | 
else  | 
| 
 
889d767ce5f4
make Minimizer honor "verbose" and "debug" options better
 
blanchet 
parents: 
41944 
diff
changeset
 | 
162  | 
""  | 
| 
41259
 
13972ced98d9
more precise error messages in "verbose" (or "debug") mode, following this morning's permission debacle
 
blanchet 
parents: 
41222 
diff
changeset
 | 
163  | 
|
| 
77430
 
51dac6fcdd0e
reverted 0506c3273814 -- the message is still useful
 
blanchet 
parents: 
77427 
diff
changeset
 | 
164  | 
fun from_lemmas [] = ""  | 
| 
 
51dac6fcdd0e
reverted 0506c3273814 -- the message is still useful
 
blanchet 
parents: 
77427 
diff
changeset
 | 
165  | 
| from_lemmas ss = " from " ^ space_implode " " (Try.serial_commas "and" (map quote ss))  | 
| 
 
51dac6fcdd0e
reverted 0506c3273814 -- the message is still useful
 
blanchet 
parents: 
77427 
diff
changeset
 | 
166  | 
|
| 74207 | 167  | 
fun string_of_atp_failure MaybeUnprovable = "Problem maybe unprovable"  | 
| 
77430
 
51dac6fcdd0e
reverted 0506c3273814 -- the message is still useful
 
blanchet 
parents: 
77427 
diff
changeset
 | 
168  | 
| string_of_atp_failure Unprovable = "Problem unprovable"  | 
| 74207 | 169  | 
| string_of_atp_failure GaveUp = "Gave up"  | 
| 75075 | 170  | 
| string_of_atp_failure ProofMissing = "Proof missing"  | 
171  | 
| string_of_atp_failure ProofIncomplete = "Proof incomplete"  | 
|
172  | 
| string_of_atp_failure ProofUnparsable = "Proof unparsable"  | 
|
| 
77430
 
51dac6fcdd0e
reverted 0506c3273814 -- the message is still useful
 
blanchet 
parents: 
77427 
diff
changeset
 | 
173  | 
| string_of_atp_failure (UnsoundProof (false, ss)) =  | 
| 
 
51dac6fcdd0e
reverted 0506c3273814 -- the message is still useful
 
blanchet 
parents: 
77427 
diff
changeset
 | 
174  | 
"Derived the lemma \"False\"" ^ from_lemmas ss ^  | 
| 
 
51dac6fcdd0e
reverted 0506c3273814 -- the message is still useful
 
blanchet 
parents: 
77427 
diff
changeset
 | 
175  | 
", likely due to the use of an unsound type encoding"  | 
| 
 
51dac6fcdd0e
reverted 0506c3273814 -- the message is still useful
 
blanchet 
parents: 
77427 
diff
changeset
 | 
176  | 
| string_of_atp_failure (UnsoundProof (true, ss)) =  | 
| 
 
51dac6fcdd0e
reverted 0506c3273814 -- the message is still useful
 
blanchet 
parents: 
77427 
diff
changeset
 | 
177  | 
"Derived the lemma \"False\"" ^ from_lemmas ss ^  | 
| 
 
51dac6fcdd0e
reverted 0506c3273814 -- the message is still useful
 
blanchet 
parents: 
77427 
diff
changeset
 | 
178  | 
", likely due to inconsistent axioms or \"sorry\"d lemmas"  | 
| 66545 | 179  | 
| string_of_atp_failure TimedOut = "Timed out"  | 
| 74207 | 180  | 
| string_of_atp_failure Inappropriate = "Problem outside the prover's scope"  | 
| 75075 | 181  | 
| string_of_atp_failure OutOfResources = "Out of resources"  | 
| 74207 | 182  | 
| string_of_atp_failure MalformedInput = "Malformed problem"  | 
183  | 
| string_of_atp_failure MalformedOutput = "Malformed output"  | 
|
184  | 
| string_of_atp_failure Interrupted = "Interrupted"  | 
|
185  | 
| string_of_atp_failure Crashed = "Crashed"  | 
|
186  | 
| string_of_atp_failure InternalError = "Internal prover error"  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
187  | 
| string_of_atp_failure (UnknownError s) =  | 
| 77272 | 188  | 
"Prover error" ^ (if s = "" then " (pass the \"verbose\" option for details)" else ":\n" ^ s)  | 
| 
39491
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
189  | 
|
| 
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
190  | 
fun extract_delimited (begin_delim, end_delim) output =  | 
| 56397 | 191  | 
(case first_field begin_delim output of  | 
192  | 
SOME (_, tail) =>  | 
|
193  | 
(case first_field "\n" tail of  | 
|
194  | 
SOME (_, tail') =>  | 
|
195  | 
if end_delim = "" then  | 
|
196  | 
tail'  | 
|
197  | 
else  | 
|
198  | 
(case first_field end_delim tail' of  | 
|
199  | 
SOME (body, _) => body  | 
|
200  | 
| NONE => "")  | 
|
201  | 
| NONE => "")  | 
|
202  | 
| NONE => "")  | 
|
| 
39491
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
203  | 
|
| 
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
204  | 
(* Splits by the first possible of a list of delimiters. *)  | 
| 
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
205  | 
fun extract_tstplike_proof delims output =  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58654 
diff
changeset
 | 
206  | 
(case apply2 (find_first (fn s => String.isSubstring s output)) (ListPair.unzip delims) of  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
207  | 
(SOME begin_delim, SOME end_delim) => extract_delimited (begin_delim, end_delim) output  | 
| 
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
208  | 
| _ => "")  | 
| 
39491
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
209  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
210  | 
fun extract_known_atp_failure known_failures output =  | 
| 
39491
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
211  | 
known_failures  | 
| 
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
212  | 
|> find_first (fn (_, pattern) => String.isSubstring pattern output)  | 
| 
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
213  | 
|> Option.map fst  | 
| 
 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 
blanchet 
parents: 
39457 
diff
changeset
 | 
214  | 
|
| 57017 | 215  | 
fun extract_tstplike_proof_and_outcome verbose proof_delims known_failures output =  | 
| 
72967
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
216  | 
let  | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
217  | 
val known_atp_failure = extract_known_atp_failure known_failures output  | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
218  | 
val tstplike_proof = extract_tstplike_proof proof_delims output  | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
219  | 
in  | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
220  | 
(case (tstplike_proof, known_atp_failure) of  | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
221  | 
      (_, SOME ProofIncomplete) => ("", NONE)
 | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
222  | 
    | (_, SOME ProofUnparsable) => ("", NONE)
 | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
223  | 
    | ("", SOME ProofMissing) => ("", NONE)
 | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
224  | 
    | ("", NONE) => ("", SOME (UnknownError (short_output verbose output)))
 | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
225  | 
    | res as ("", _) => res
 | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
226  | 
| (tstplike_proof, _) => (tstplike_proof, NONE))  | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
227  | 
end  | 
| 39452 | 228  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
53225 
diff
changeset
 | 
229  | 
type atp_step_name = string * string list  | 
| 39452 | 230  | 
|
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42966 
diff
changeset
 | 
231  | 
fun is_same_atp_step (s1, _) (s2, _) = s1 = s2  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42966 
diff
changeset
 | 
232  | 
|
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42966 
diff
changeset
 | 
233  | 
val vampire_fact_prefix = "f"  | 
| 39452 | 234  | 
|
| 
52755
 
4183c3219745
simplified Vampire hack -- no need to run it for other ATPs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
235  | 
fun vampire_step_name_ord p =  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58654 
diff
changeset
 | 
236  | 
let val q = apply2 fst p in  | 
| 
75058
 
14e27dedee10
made sorting of Vampire facts more robust in the face of names that deviate from the standard scheme
 
blanchet 
parents: 
75055 
diff
changeset
 | 
237  | 
(* The "unprefix" part is to cope with Vampire's output, which puts facts with names of the  | 
| 
 
14e27dedee10
made sorting of Vampire facts more robust in the face of names that deviate from the standard scheme
 
blanchet 
parents: 
75055 
diff
changeset
 | 
238  | 
form "fN" where N is an integer in reverse order. *)  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58654 
diff
changeset
 | 
239  | 
(case apply2 (Int.fromString o perhaps (try (unprefix vampire_fact_prefix))) q of  | 
| 
52755
 
4183c3219745
simplified Vampire hack -- no need to run it for other ATPs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
240  | 
(SOME i, SOME j) => int_ord (i, j)  | 
| 
75058
 
14e27dedee10
made sorting of Vampire facts more robust in the face of names that deviate from the standard scheme
 
blanchet 
parents: 
75055 
diff
changeset
 | 
241  | 
| (SOME _, NONE) => LESS  | 
| 
 
14e27dedee10
made sorting of Vampire facts more robust in the face of names that deviate from the standard scheme
 
blanchet 
parents: 
75055 
diff
changeset
 | 
242  | 
| (NONE, SOME _) => GREATER  | 
| 
 
14e27dedee10
made sorting of Vampire facts more robust in the face of names that deviate from the standard scheme
 
blanchet 
parents: 
75055 
diff
changeset
 | 
243  | 
| (NONE, NONE) => string_ord q)  | 
| 39452 | 244  | 
end  | 
245  | 
||
| 54811 | 246  | 
type ('a, 'b) atp_step = atp_step_name * atp_formula_role * 'a * 'b * atp_step_name list
 | 
| 39452 | 247  | 
|
| 54811 | 248  | 
type 'a atp_proof = (('a, 'a, ('a, 'a atp_type) atp_term, 'a) atp_formula, string) atp_step list
 | 
| 39452 | 249  | 
|
250  | 
(**** PARSING OF TSTP FORMAT ****)  | 
|
251  | 
||
| 
56404
 
9cb137ec6ec8
use Z3 TPTP cores rather than proofs since the latter are somewhat broken
 
blanchet 
parents: 
56397 
diff
changeset
 | 
252  | 
(* Strings enclosed in single quotes (e.g., file names), identifiers possibly starting  | 
| 75038 | 253  | 
with "$" and possibly with "!" in them. *)  | 
| 39452 | 254  | 
val scan_general_id =  | 
| 
47917
 
b287682bf917
improve parsing of Waldmeister dependencies (and kill obsolete hack)
 
blanchet 
parents: 
47787 
diff
changeset
 | 
255  | 
$$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58654 
diff
changeset
 | 
256  | 
|| (Scan.repeat ($$ "$") -- Scan.many1 Symbol.is_letdig >> (op ^ o apply2 implode))  | 
| 
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58654 
diff
changeset
 | 
257  | 
-- Scan.optional (Scan.repeat ($$ "!") -- Scan.many1 Symbol.is_letdig >> (op ^ o apply2 implode)) ""  | 
| 
56404
 
9cb137ec6ec8
use Z3 TPTP cores rather than proofs since the latter are somewhat broken
 
blanchet 
parents: 
56397 
diff
changeset
 | 
258  | 
>> op ^  | 
| 39452 | 259  | 
|
| 70805 | 260  | 
fun skip_term x =  | 
| 
45208
 
9a00f9cc8707
marginally cleaner proof parsing, that doesn't stumble upon LEO-II's E-step proofs
 
blanchet 
parents: 
45203 
diff
changeset
 | 
261  | 
let  | 
| 
45235
 
7187bce94e88
more robust parsing of TSTP sources -- Vampire has nonstandard "introduced()" tags and Waldmeister(OnTPTP) has weird "theory(...)" dependencies
 
blanchet 
parents: 
45209 
diff
changeset
 | 
262  | 
fun skip _ accum [] = (accum, [])  | 
| 
54803
 
41a109a00c53
fixed variable confusion introduced by 'tuning' change 565f9af86d67
 
blanchet 
parents: 
54799 
diff
changeset
 | 
263  | 
| skip n accum (ss as s :: ss') =  | 
| 
57707
 
0242e9578828
imported patch satallax_proof_support_Sledgehammer
 
fleury 
parents: 
57697 
diff
changeset
 | 
264  | 
if (s = "," orelse s = ".") andalso n = 0 then  | 
| 54799 | 265  | 
(accum, ss)  | 
| 56397 | 266  | 
else if member (op =) [")", "]"] s then  | 
| 
54803
 
41a109a00c53
fixed variable confusion introduced by 'tuning' change 565f9af86d67
 
blanchet 
parents: 
54799 
diff
changeset
 | 
267  | 
if n = 0 then (accum, ss) else skip (n - 1) (s :: accum) ss'  | 
| 56397 | 268  | 
        else if member (op =) ["(", "["] s then
 | 
| 
54803
 
41a109a00c53
fixed variable confusion introduced by 'tuning' change 565f9af86d67
 
blanchet 
parents: 
54799 
diff
changeset
 | 
269  | 
skip (n + 1) (s :: accum) ss'  | 
| 54799 | 270  | 
else  | 
| 
54803
 
41a109a00c53
fixed variable confusion introduced by 'tuning' change 565f9af86d67
 
blanchet 
parents: 
54799 
diff
changeset
 | 
271  | 
skip n (s :: accum) ss'  | 
| 54799 | 272  | 
in  | 
| 70805 | 273  | 
(skip 0 [] #>> (rev #> implode)) x  | 
| 54799 | 274  | 
end  | 
| 70805 | 275  | 
and skip_terms x = (skip_term ::: Scan.repeat ($$ "," |-- skip_term)) x  | 
| 
45208
 
9a00f9cc8707
marginally cleaner proof parsing, that doesn't stumble upon LEO-II's E-step proofs
 
blanchet 
parents: 
45203 
diff
changeset
 | 
276  | 
|
| 
 
9a00f9cc8707
marginally cleaner proof parsing, that doesn't stumble upon LEO-II's E-step proofs
 
blanchet 
parents: 
45203 
diff
changeset
 | 
277  | 
datatype source =  | 
| 75122 | 278  | 
File_Source of string * string option  | 
279  | 
| Inference_Source of string * string list  | 
|
280  | 
| Introduced_Source of string  | 
|
281  | 
| Define_Source  | 
|
| 
45208
 
9a00f9cc8707
marginally cleaner proof parsing, that doesn't stumble upon LEO-II's E-step proofs
 
blanchet 
parents: 
45203 
diff
changeset
 | 
282  | 
|
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48130 
diff
changeset
 | 
283  | 
val dummy_phi = AAtom (ATerm (("", []), []))
 | 
| 
57697
 
44341963ade3
correctly translate THF functions from terms to types
 
blanchet 
parents: 
57656 
diff
changeset
 | 
284  | 
val dummy_atype = AType (("", []), [])
 | 
| 
45235
 
7187bce94e88
more robust parsing of TSTP sources -- Vampire has nonstandard "introduced()" tags and Waldmeister(OnTPTP) has weird "theory(...)" dependencies
 
blanchet 
parents: 
45209 
diff
changeset
 | 
285  | 
|
| 
54769
 
3d6ac2f68bf3
correcly recognize E skolemization steps that are wrapped in a 'shift_quantors' inference
 
blanchet 
parents: 
53587 
diff
changeset
 | 
286  | 
(* "skip_term" is there to cope with Waldmeister nonsense such as "theory(equality)". *)  | 
| 50011 | 287  | 
fun parse_dependency x =  | 
288  | 
(parse_inference_source >> snd  | 
|
289  | 
|| scan_general_id --| skip_term >> single) x  | 
|
290  | 
and parse_dependencies x =  | 
|
| 61476 | 291  | 
(Scan.repeats (Scan.option ($$ ",") |-- parse_dependency)  | 
292  | 
>> (filter_out (curry (op =) "theory"))) x  | 
|
| 50011 | 293  | 
and parse_file_source x =  | 
294  | 
  (Scan.this_string "file" |-- $$ "(" |-- scan_general_id
 | 
|
| 57154 | 295  | 
-- Scan.option ($$ "," |-- scan_general_id  | 
| 
57266
 
6a3b5085fb8f
fixed parsing of one-argument 'file()' in TSTP files
 
blanchet 
parents: 
57261 
diff
changeset
 | 
296  | 
--| Scan.option ($$ "," |-- $$ "[" -- Scan.option scan_general_id --| $$ "]")) --| $$ ")") x  | 
| 50011 | 297  | 
and parse_inference_source x =  | 
298  | 
  (Scan.this_string "inference" |-- $$ "(" |-- scan_general_id
 | 
|
299  | 
--| skip_term --| $$ "," --| skip_term --| $$ "," --| $$ "["  | 
|
300  | 
-- parse_dependencies --| $$ "]" --| $$ ")") x  | 
|
| 57785 | 301  | 
and parse_introduced_source x =  | 
302  | 
  (Scan.this_string "introduced" |-- $$ "(" |-- scan_general_id
 | 
|
303  | 
--| Scan.option ($$ "," |-- skip_term) --| $$ ")") x  | 
|
| 75122 | 304  | 
and parse_define_source x =  | 
305  | 
  (Scan.this_string "define" |-- $$ "(" |-- skip_term --| $$ ")") x
 | 
|
| 50011 | 306  | 
and parse_source x =  | 
| 
72967
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
307  | 
(parse_file_source >> File_Source >> SOME  | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
308  | 
|| parse_inference_source >> Inference_Source >> SOME  | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
309  | 
|| parse_introduced_source >> Introduced_Source >> SOME  | 
| 75122 | 310  | 
|| parse_define_source >> K Define_Source >> SOME  | 
| 
72967
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
311  | 
   || scan_general_id >> (fn s => SOME (Inference_Source ("", [s]))) (* for E *)
 | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
312  | 
|| skip_term >> K NONE) x  | 
| 39452 | 313  | 
|
| 54799 | 314  | 
fun list_app (f, args) = fold (fn arg => fn f => ATerm ((tptp_app, []), [f, arg])) args f  | 
| 
42966
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42965 
diff
changeset
 | 
315  | 
|
| 54820 | 316  | 
fun parse_class x = scan_general_id x  | 
317  | 
and parse_classes x = (parse_class ::: Scan.repeat ($$ "&" |-- parse_class)) x  | 
|
| 54811 | 318  | 
|
319  | 
fun parse_type x =  | 
|
| 66428 | 320  | 
  (($$ "(" |-- parse_type --| $$ ")"
 | 
| 70805 | 321  | 
|| Scan.this_string tptp_pi_binder |-- $$ "[" |-- skip_terms --| $$ "]" --| $$ ":" -- parse_type  | 
| 
67022
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
322  | 
>> (fn (_, ty) => ty (* currently ignoring type constructor declarations anyway *))  | 
| 66428 | 323  | 
    || (scan_general_id -- Scan.optional ($$ "{" |-- parse_classes --| $$ "}") [])
 | 
324  | 
        -- Scan.optional ($$ "(" |-- parse_types --| $$ ")") []
 | 
|
325  | 
>> AType)  | 
|
| 
67022
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
326  | 
-- Scan.option (($$ tptp_app || $$ tptp_fun_type || $$ tptp_product_type) -- parse_type)  | 
| 66428 | 327  | 
>> (fn (a, NONE) => a  | 
| 
67022
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
328  | 
| (a, SOME (bin_op, b)) =>  | 
| 
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
329  | 
if bin_op = tptp_app then  | 
| 
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
330  | 
(case a of  | 
| 
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
331  | 
AType (s_clss, tys) => AType (s_clss, tys @ [b])  | 
| 
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
332  | 
| _ => raise UNRECOGNIZED_ATP_PROOF ())  | 
| 
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
333  | 
else if bin_op = tptp_fun_type then  | 
| 
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
334  | 
AFun (a, b)  | 
| 
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
335  | 
else if bin_op = tptp_product_type then  | 
| 
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
336  | 
AType ((tptp_product_type, []), [a, b])  | 
| 
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
337  | 
else  | 
| 
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
338  | 
raise Fail "impossible case")) x  | 
| 66544 | 339  | 
and parse_types x =  | 
340  | 
(parse_type ::: Scan.repeat ($$ "," |-- parse_type)) x  | 
|
| 54811 | 341  | 
|
| 58477 | 342  | 
(* We currently half ignore types. *)  | 
| 72398 | 343  | 
fun parse_fol_optional_type_signature x =  | 
| 58600 | 344  | 
(Scan.option ($$ tptp_has_type |-- parse_type)  | 
345  | 
>> (fn some as SOME (AType ((s, []), [])) => if s = dfg_individual_type then NONE else some  | 
|
346  | 
| res => res)) x  | 
|
| 72398 | 347  | 
and parse_fol_arg x =  | 
348  | 
  ($$ "(" |-- parse_fol_term --| $$ ")" --| parse_fol_optional_type_signature
 | 
|
349  | 
|| scan_general_id -- parse_fol_optional_type_signature  | 
|
| 54811 | 350  | 
-- Scan.optional ($$ "<" |-- parse_types --| $$ ">") []  | 
| 72398 | 351  | 
       -- Scan.optional ($$ "(" |-- parse_fol_terms --| $$ ")") []
 | 
| 58477 | 352  | 
>> (fn (((s, ty_opt), tyargs), args) =>  | 
353  | 
if is_tptp_variable s andalso null tyargs andalso null args andalso is_some ty_opt then  | 
|
354  | 
ATerm ((s, the_list ty_opt), [])  | 
|
355  | 
else  | 
|
356  | 
ATerm ((s, tyargs), args))) x  | 
|
| 72398 | 357  | 
and parse_fol_term x =  | 
358  | 
(parse_fol_arg -- Scan.repeat ($$ tptp_app |-- parse_fol_arg)  | 
|
359  | 
--| parse_fol_optional_type_signature >> list_app) x  | 
|
360  | 
and parse_fol_terms x = (parse_fol_term ::: Scan.repeat ($$ "," |-- parse_fol_term)) x  | 
|
| 39452 | 361  | 
|
| 72398 | 362  | 
fun parse_fol_atom x =  | 
363  | 
(parse_fol_term --  | 
|
364  | 
Scan.option (Scan.option ($$ tptp_not_infix) --| $$ tptp_equal -- parse_fol_term)  | 
|
| 39598 | 365  | 
>> (fn (u1, NONE) => AAtom u1  | 
| 45881 | 366  | 
| (u1, SOME (neg, u2)) =>  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48130 
diff
changeset
 | 
367  | 
          AAtom (ATerm (("equal", []), [u1, u2])) |> is_some neg ? mk_anot)) x
 | 
| 39452 | 368  | 
|
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
369  | 
(* TPTP formulas are fully parenthesized, so we don't need to worry about operator precedence. *)  | 
| 72398 | 370  | 
fun parse_fol_literal x =  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42966 
diff
changeset
 | 
371  | 
((Scan.repeat ($$ tptp_not) >> length)  | 
| 72398 | 372  | 
      -- ($$ "(" |-- parse_fol_formula --| $$ ")"
 | 
373  | 
|| parse_fol_quantified_formula  | 
|
374  | 
|| parse_fol_atom)  | 
|
| 
42605
 
8734eb0033b3
Vampire sometimes generates formulas with ~ (not) followed by a quantified subformula, without parentheses -- parse these correctly
 
blanchet 
parents: 
42603 
diff
changeset
 | 
375  | 
>> (fn (n, phi) => phi |> n mod 2 = 1 ? mk_anot)) x  | 
| 72398 | 376  | 
and parse_fol_formula x =  | 
377  | 
(parse_fol_literal  | 
|
| 43163 | 378  | 
-- Scan.option ((Scan.this_string tptp_implies  | 
379  | 
|| Scan.this_string tptp_iff  | 
|
380  | 
|| Scan.this_string tptp_not_iff  | 
|
381  | 
|| Scan.this_string tptp_if  | 
|
382  | 
|| $$ tptp_or  | 
|
| 72398 | 383  | 
|| $$ tptp_and) -- parse_fol_formula)  | 
| 39452 | 384  | 
>> (fn (phi1, NONE) => phi1  | 
| 43163 | 385  | 
| (phi1, SOME (c, phi2)) =>  | 
386  | 
if c = tptp_implies then mk_aconn AImplies phi1 phi2  | 
|
387  | 
else if c = tptp_iff then mk_aconn AIff phi1 phi2  | 
|
388  | 
else if c = tptp_not_iff then mk_anot (mk_aconn AIff phi1 phi2)  | 
|
389  | 
else if c = tptp_if then mk_aconn AImplies phi2 phi1  | 
|
390  | 
else if c = tptp_or then mk_aconn AOr phi1 phi2  | 
|
391  | 
else if c = tptp_and then mk_aconn AAnd phi1 phi2  | 
|
392  | 
          else raise Fail ("impossible connective " ^ quote c))) x
 | 
|
| 72398 | 393  | 
and parse_fol_quantified_formula x =  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42966 
diff
changeset
 | 
394  | 
(($$ tptp_forall >> K AForall || $$ tptp_exists >> K AExists)  | 
| 72398 | 395  | 
--| $$ "[" -- parse_fol_terms --| $$ "]" --| $$ ":" -- parse_fol_literal  | 
| 61030 | 396  | 
>> (fn ((q, ts), phi) => AQuant (q, map (fn ATerm ((s, _), _) => (s, NONE)) ts, phi))) x  | 
| 39452 | 397  | 
|
398  | 
val parse_tstp_extra_arguments =  | 
|
| 
72967
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
399  | 
Scan.optional ($$ "," |-- parse_source --| Scan.option ($$ "," |-- skip_term)) NONE  | 
| 39452 | 400  | 
|
| 
47927
 
c35238d19bb9
repair the Waldmeister endgame only for Waldmeister proofs
 
blanchet 
parents: 
47926 
diff
changeset
 | 
401  | 
val waldmeister_conjecture_name = "conjecture_1"  | 
| 42943 | 402  | 
|
403  | 
fun is_same_term subst tm1 tm2 =  | 
|
404  | 
let  | 
|
| 57257 | 405  | 
fun do_term_pair (AAbs (((var1, typ1), body1), args1)) (AAbs (((var2, typ2), body2), args2))  | 
406  | 
(SOME subst) =  | 
|
407  | 
if typ1 <> typ2 andalso length args1 = length args2 then NONE  | 
|
408  | 
else  | 
|
409  | 
let val ls = length subst in  | 
|
410  | 
SOME ((var1, var2) :: subst)  | 
|
411  | 
|> do_term_pair body1 body2  | 
|
412  | 
|> (fn SOME subst => SOME (nth_drop (length subst - ls - 1) subst)  | 
|
413  | 
| NONE => NONE)  | 
|
414  | 
|> (if length args1 = length args2  | 
|
415  | 
then fold2 do_term_pair args1 args2  | 
|
| 
57258
 
67d85a8aa6cc
Moving the remote prefix deleting on Sledgehammer's side
 
fleury 
parents: 
57257 
diff
changeset
 | 
416  | 
else K NONE)  | 
| 57257 | 417  | 
end  | 
418  | 
| do_term_pair (ATerm ((s1, _), args1)) (ATerm ((s2, _), args2)) (SOME subst) =  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58654 
diff
changeset
 | 
419  | 
(case apply2 is_tptp_variable (s1, s2) of  | 
| 42943 | 420  | 
(true, true) =>  | 
421  | 
(case AList.lookup (op =) subst s1 of  | 
|
| 57257 | 422  | 
SOME s2' => if s2' = s2 then SOME subst else NONE  | 
423  | 
| NONE =>  | 
|
424  | 
if null (AList.find (op =) subst s2) then SOME ((s1, s2) :: subst)  | 
|
425  | 
else NONE)  | 
|
| 42943 | 426  | 
| (false, false) =>  | 
| 57257 | 427  | 
if s1 = s2 then  | 
428  | 
SOME subst  | 
|
| 42943 | 429  | 
else  | 
430  | 
NONE  | 
|
| 57257 | 431  | 
| _ => NONE) |> (if length args1 = length args2  | 
432  | 
then fold2 do_term_pair args1 args2  | 
|
| 
57258
 
67d85a8aa6cc
Moving the remote prefix deleting on Sledgehammer's side
 
fleury 
parents: 
57257 
diff
changeset
 | 
433  | 
else K NONE)  | 
| 57257 | 434  | 
| do_term_pair _ _ _ = NONE  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
435  | 
in  | 
| 57257 | 436  | 
SOME subst |> do_term_pair tm1 tm2 |> is_some  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
437  | 
end  | 
| 42943 | 438  | 
|
| 
47921
 
fc26d5538868
ensure the "show" equation is not reoriented by Waldmeister
 
blanchet 
parents: 
47917 
diff
changeset
 | 
439  | 
fun is_same_formula comm subst (AQuant (q1, xs1, phi1)) (AQuant (q2, xs2, phi2)) =  | 
| 42943 | 440  | 
q1 = q2 andalso length xs1 = length xs2 andalso  | 
| 77917 | 441  | 
is_same_formula comm (map2 (fn (x1, _) => fn (x2, _) => (x1, x2)) xs1 xs2 @ subst) phi1 phi2  | 
| 
47921
 
fc26d5538868
ensure the "show" equation is not reoriented by Waldmeister
 
blanchet 
parents: 
47917 
diff
changeset
 | 
442  | 
| is_same_formula comm subst (AConn (c1, phis1)) (AConn (c2, phis2)) =  | 
| 
77918
 
55b81d14a1b8
tuned; avoided intermediate list and list traversal
 
desharna 
parents: 
77917 
diff
changeset
 | 
443  | 
c1 = c2 andalso forall2 (is_same_formula comm subst) phis1 phis2  | 
| 57202 | 444  | 
  | is_same_formula comm subst (AAtom (tm1 as ATerm (("equal", tys), [tm11, tm12]))) (AAtom tm2) =
 | 
| 47926 | 445  | 
is_same_term subst tm1 tm2 orelse  | 
| 54799 | 446  | 
    (comm andalso is_same_term subst (ATerm (("equal", tys), [tm12, tm11])) tm2)
 | 
| 
47921
 
fc26d5538868
ensure the "show" equation is not reoriented by Waldmeister
 
blanchet 
parents: 
47917 
diff
changeset
 | 
447  | 
| is_same_formula _ subst (AAtom tm1) (AAtom tm2) = is_same_term subst tm1 tm2  | 
| 
 
fc26d5538868
ensure the "show" equation is not reoriented by Waldmeister
 
blanchet 
parents: 
47917 
diff
changeset
 | 
448  | 
| is_same_formula _ _ _ _ = false  | 
| 42943 | 449  | 
|
| 
50521
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
50236 
diff
changeset
 | 
450  | 
fun matching_formula_line_identifier phi (Formula ((ident, _), _, phi', _, _)) =  | 
| 
47921
 
fc26d5538868
ensure the "show" equation is not reoriented by Waldmeister
 
blanchet 
parents: 
47917 
diff
changeset
 | 
451  | 
if is_same_formula true [] phi phi' then SOME (ident, phi') else NONE  | 
| 42943 | 452  | 
| matching_formula_line_identifier _ _ = NONE  | 
453  | 
||
| 54799 | 454  | 
fun find_formula_in_problem phi =  | 
455  | 
maps snd  | 
|
456  | 
#> map_filter (matching_formula_line_identifier phi)  | 
|
457  | 
#> try (single o hd)  | 
|
458  | 
#> the_default []  | 
|
| 42943 | 459  | 
|
| 54799 | 460  | 
fun commute_eq (AAtom (ATerm ((s, tys), tms))) = AAtom (ATerm ((s, tys), rev tms))  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47927 
diff
changeset
 | 
461  | 
| commute_eq _ = raise Fail "expected equation"  | 
| 
47921
 
fc26d5538868
ensure the "show" equation is not reoriented by Waldmeister
 
blanchet 
parents: 
47917 
diff
changeset
 | 
462  | 
|
| 
50012
 
01cb92151a53
track formula roles in proofs and use that to determine whether the conjecture should be negated or not
 
blanchet 
parents: 
50011 
diff
changeset
 | 
463  | 
fun role_of_tptp_string "axiom" = Axiom  | 
| 
 
01cb92151a53
track formula roles in proofs and use that to determine whether the conjecture should be negated or not
 
blanchet 
parents: 
50011 
diff
changeset
 | 
464  | 
| role_of_tptp_string "definition" = Definition  | 
| 
 
01cb92151a53
track formula roles in proofs and use that to determine whether the conjecture should be negated or not
 
blanchet 
parents: 
50011 
diff
changeset
 | 
465  | 
| role_of_tptp_string "lemma" = Lemma  | 
| 
 
01cb92151a53
track formula roles in proofs and use that to determine whether the conjecture should be negated or not
 
blanchet 
parents: 
50011 
diff
changeset
 | 
466  | 
| role_of_tptp_string "hypothesis" = Hypothesis  | 
| 
 
01cb92151a53
track formula roles in proofs and use that to determine whether the conjecture should be negated or not
 
blanchet 
parents: 
50011 
diff
changeset
 | 
467  | 
| role_of_tptp_string "conjecture" = Conjecture  | 
| 
 
01cb92151a53
track formula roles in proofs and use that to determine whether the conjecture should be negated or not
 
blanchet 
parents: 
50011 
diff
changeset
 | 
468  | 
| role_of_tptp_string "negated_conjecture" = Negated_Conjecture  | 
| 
 
01cb92151a53
track formula roles in proofs and use that to determine whether the conjecture should be negated or not
 
blanchet 
parents: 
50011 
diff
changeset
 | 
469  | 
| role_of_tptp_string "plain" = Plain  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
470  | 
| role_of_tptp_string "type" = Type_Role  | 
| 
50012
 
01cb92151a53
track formula roles in proofs and use that to determine whether the conjecture should be negated or not
 
blanchet 
parents: 
50011 
diff
changeset
 | 
471  | 
| role_of_tptp_string _ = Unknown  | 
| 
 
01cb92151a53
track formula roles in proofs and use that to determine whether the conjecture should be negated or not
 
blanchet 
parents: 
50011 
diff
changeset
 | 
472  | 
|
| 
57714
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57709 
diff
changeset
 | 
473  | 
fun parse_one_in_list xs =  | 
| 
 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 
fleury 
parents: 
57709 
diff
changeset
 | 
474  | 
foldl1 (op ||) (map Scan.this_string xs)  | 
| 57709 | 475  | 
|
| 
75055
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
476  | 
val tptp_literal_binary_ops = [tptp_equal, tptp_not_equal]  | 
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
477  | 
val tptp_nonliteral_binary_ops =  | 
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
478  | 
[tptp_and, tptp_or, tptp_implies, tptp_iff, tptp_if, tptp_not_and, tptp_not_or, tptp_not_iff]  | 
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
479  | 
|
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
480  | 
fun parse_literal_binary_op x =  | 
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
481  | 
(parse_one_in_list tptp_literal_binary_ops  | 
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
482  | 
>> (fn c => if c = tptp_equal then "equal" else c)) x  | 
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
483  | 
|
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
484  | 
fun parse_nonliteral_binary_op x =  | 
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
485  | 
(parse_one_in_list tptp_nonliteral_binary_ops  | 
| 
57697
 
44341963ade3
correctly translate THF functions from terms to types
 
blanchet 
parents: 
57656 
diff
changeset
 | 
486  | 
>> (fn c => if c = tptp_equal then "equal" else c)) x  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
487  | 
|
| 72398 | 488  | 
val parse_fol_quantifier =  | 
| 57709 | 489  | 
parse_one_in_list [tptp_forall, tptp_exists, tptp_lambda, tptp_hilbert_choice, tptp_hilbert_the]  | 
490  | 
||
| 72398 | 491  | 
val parse_hol_quantifier =  | 
| 57709 | 492  | 
parse_one_in_list [tptp_ho_forall, tptp_ho_exists, tptp_hilbert_choice, tptp_hilbert_the]  | 
493  | 
||
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
494  | 
fun mk_ho_of_fo_quant q =  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
495  | 
if q = tptp_forall then tptp_ho_forall  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
496  | 
else if q = tptp_exists then tptp_ho_exists  | 
| 57709 | 497  | 
else if q = tptp_hilbert_choice then tptp_hilbert_choice  | 
498  | 
else if q = tptp_hilbert_the then tptp_hilbert_the  | 
|
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
499  | 
  else raise Fail ("unrecognized quantification: " ^ q)
 | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
500  | 
|
| 74393 | 501  | 
fun remove_hol_app (ATerm ((s, ty), args)) =  | 
| 
74049
 
d0b190b4f15d
parse TPTP operator @ also when not parenthesized
 
blanchet 
parents: 
74006 
diff
changeset
 | 
502  | 
if s = tptp_app then  | 
| 74393 | 503  | 
(case args of  | 
504  | 
ATerm (f, xs) :: ys => remove_hol_app (ATerm (f, xs @ ys))  | 
|
505  | 
| AAbs ((var, phi), xs) :: ys => remove_hol_app (AAbs ((var, phi), xs @ ys)))  | 
|
| 
61329
 
426c9c858099
added "!=" (disequality) as a TPTP binary operator, since it pops up in LEO-II proofs
 
blanchet 
parents: 
61030 
diff
changeset
 | 
506  | 
else  | 
| 74393 | 507  | 
ATerm ((s, ty), map remove_hol_app args)  | 
508  | 
| remove_hol_app (AAbs ((var, phi), args)) =  | 
|
509  | 
AAbs ((var, remove_hol_app phi), map remove_hol_app args)  | 
|
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
510  | 
|
| 72398 | 511  | 
fun parse_hol_typed_var x =  | 
| 66428 | 512  | 
(Scan.repeat (scan_general_id -- Scan.option ($$ tptp_has_type |-- parse_type)  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
513  | 
--| Scan.option (Scan.this_string ","))  | 
| 72398 | 514  | 
   || $$ "(" |-- parse_hol_typed_var --| $$ ")") x
 | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
515  | 
|
| 72398 | 516  | 
fun parse_simple_hol_term x =  | 
| 
78696
 
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
 
blanchet 
parents: 
78694 
diff
changeset
 | 
517  | 
(parse_fol_quantifier -- ($$ "[" |-- parse_hol_typed_var --| $$ "]" --| $$ ":")  | 
| 
 
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
 
blanchet 
parents: 
78694 
diff
changeset
 | 
518  | 
-- parse_simple_hol_term  | 
| 
 
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
 
blanchet 
parents: 
78694 
diff
changeset
 | 
519  | 
>> (fn ((q, ys), t) =>  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
520  | 
fold_rev  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
521  | 
(fn (var, ty) => fn r =>  | 
| 57257 | 522  | 
AAbs (((var, the_default dummy_atype ty), r), [])  | 
523  | 
|> (if tptp_lambda <> q then  | 
|
| 
75055
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
524  | 
mk_app (q |> mk_ho_of_fo_quant |> mk_simple_aterm)  | 
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
525  | 
else  | 
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
526  | 
I))  | 
| 57257 | 527  | 
ys t)  | 
| 
78691
 
1320782a394e
improved Sledgehammer's HOL proof parser w.r.t. negation
 
blanchet 
parents: 
77918 
diff
changeset
 | 
528  | 
|| Scan.this_string tptp_not |-- parse_simple_hol_term >> mk_app (mk_simple_aterm tptp_not)  | 
| 66428 | 529  | 
|| scan_general_id -- Scan.option ($$ tptp_has_type |-- parse_type)  | 
| 
61329
 
426c9c858099
added "!=" (disequality) as a TPTP binary operator, since it pops up in LEO-II proofs
 
blanchet 
parents: 
61030 
diff
changeset
 | 
530  | 
>> (fn (var, typ_opt) => ATerm ((var, the_list typ_opt), []))  | 
| 72398 | 531  | 
|| parse_hol_quantifier >> mk_simple_aterm  | 
532  | 
  || $$ "(" |-- parse_hol_term --| $$ ")"
 | 
|
| 
78694
 
5e995ceb7490
allow (~) syntax in TPTP proofs for unapplied negation
 
blanchet 
parents: 
78692 
diff
changeset
 | 
533  | 
|| Scan.this_string tptp_not >> mk_simple_aterm  | 
| 
75055
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
534  | 
|| parse_literal_binary_op >> mk_simple_aterm  | 
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
535  | 
|| parse_nonliteral_binary_op >> mk_simple_aterm) x  | 
| 
75050
 
632c9ae415fa
adjust TPTP THF parser to give priority to @ over other operators, to parse Ehoh proofs
 
blanchet 
parents: 
75038 
diff
changeset
 | 
536  | 
and parse_applied_hol_term x =  | 
| 
 
632c9ae415fa
adjust TPTP THF parser to give priority to @ over other operators, to parse Ehoh proofs
 
blanchet 
parents: 
75038 
diff
changeset
 | 
537  | 
(parse_simple_hol_term -- Scan.repeat (Scan.this_string tptp_app |-- parse_simple_hol_term)  | 
| 
 
632c9ae415fa
adjust TPTP THF parser to give priority to @ over other operators, to parse Ehoh proofs
 
blanchet 
parents: 
75038 
diff
changeset
 | 
538  | 
>> (fn (t1, tis) => fold (fn ti => fn left => mk_app left ti) tis t1)) x  | 
| 
75055
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
539  | 
and parse_literal_hol_term x =  | 
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
540  | 
(parse_applied_hol_term -- Scan.repeat (parse_literal_binary_op -- parse_applied_hol_term)  | 
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
541  | 
>> (fn (t1, c_ti_s) =>  | 
| 
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
542  | 
fold (fn (c, ti) => fn left => mk_apps (mk_simple_aterm c) [left, ti]) c_ti_s t1)) x  | 
| 72398 | 543  | 
and parse_hol_term x =  | 
| 
75055
 
c84a20e9b00f
correctly parse E proofs that assume '=' and '!=' bind more tightly than connectives
 
blanchet 
parents: 
75050 
diff
changeset
 | 
544  | 
(parse_literal_hol_term -- Scan.repeat (parse_nonliteral_binary_op -- parse_literal_hol_term)  | 
| 
74049
 
d0b190b4f15d
parse TPTP operator @ also when not parenthesized
 
blanchet 
parents: 
74006 
diff
changeset
 | 
545  | 
>> (fn (t1, c_ti_s) =>  | 
| 
78692
 
1b0f5576f5e9
use same associativity as Isabelle when parsing HOL proofs
 
blanchet 
parents: 
78691 
diff
changeset
 | 
546  | 
let  | 
| 
 
1b0f5576f5e9
use same associativity as Isabelle when parsing HOL proofs
 
blanchet 
parents: 
78691 
diff
changeset
 | 
547  | 
val cs = map fst c_ti_s  | 
| 
 
1b0f5576f5e9
use same associativity as Isabelle when parsing HOL proofs
 
blanchet 
parents: 
78691 
diff
changeset
 | 
548  | 
val tis = t1 :: map snd c_ti_s  | 
| 
 
1b0f5576f5e9
use same associativity as Isabelle when parsing HOL proofs
 
blanchet 
parents: 
78691 
diff
changeset
 | 
549  | 
val (tis_but_k, tk) = split_last tis  | 
| 
 
1b0f5576f5e9
use same associativity as Isabelle when parsing HOL proofs
 
blanchet 
parents: 
78691 
diff
changeset
 | 
550  | 
in  | 
| 
 
1b0f5576f5e9
use same associativity as Isabelle when parsing HOL proofs
 
blanchet 
parents: 
78691 
diff
changeset
 | 
551  | 
fold_rev (fn (ti, c) => fn right => mk_apps (mk_simple_aterm c) [ti, right])  | 
| 
 
1b0f5576f5e9
use same associativity as Isabelle when parsing HOL proofs
 
blanchet 
parents: 
78691 
diff
changeset
 | 
552  | 
(tis_but_k ~~ cs) tk  | 
| 
 
1b0f5576f5e9
use same associativity as Isabelle when parsing HOL proofs
 
blanchet 
parents: 
78691 
diff
changeset
 | 
553  | 
end)) x  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
554  | 
|
| 72398 | 555  | 
fun parse_hol_formula x = (parse_hol_term #>> remove_hol_app #>> AAtom) x  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
556  | 
|
| 75064 | 557  | 
fun parse_tstp_hol_line full problem =  | 
| 
61329
 
426c9c858099
added "!=" (disequality) as a TPTP binary operator, since it pops up in LEO-II proofs
 
blanchet 
parents: 
61030 
diff
changeset
 | 
558  | 
  (Scan.this_string tptp_thf -- $$ "(") |-- scan_general_id --| $$ ","
 | 
| 75064 | 559  | 
-- Symbol.scan_ascii_id --| $$ "," --  | 
560  | 
(if full then parse_hol_formula || skip_term >> K dummy_phi else skip_term >> K dummy_phi)  | 
|
| 
67022
 
49309fe530fd
more robust parsing for THF proofs (esp. polymorphic Leo-III proofs)
 
blanchet 
parents: 
67021 
diff
changeset
 | 
561  | 
-- parse_tstp_extra_arguments --| $$ ")"  | 
| 
61329
 
426c9c858099
added "!=" (disequality) as a TPTP binary operator, since it pops up in LEO-II proofs
 
blanchet 
parents: 
61030 
diff
changeset
 | 
562  | 
--| $$ "."  | 
| 
75123
 
66eb6fdfc244
handle Zipperposition definitions in Isar proof construction
 
blanchet 
parents: 
75122 
diff
changeset
 | 
563  | 
>> (fn (((num, role0), phi), src) =>  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
564  | 
let  | 
| 
75123
 
66eb6fdfc244
handle Zipperposition definitions in Isar proof construction
 
blanchet 
parents: 
75122 
diff
changeset
 | 
565  | 
val role = role_of_tptp_string role0  | 
| 
 
66eb6fdfc244
handle Zipperposition definitions in Isar proof construction
 
blanchet 
parents: 
75122 
diff
changeset
 | 
566  | 
val (name, phi, role', rule, deps) =  | 
| 
72967
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
567  | 
(case src of  | 
| 
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
568  | 
SOME (File_Source (_, SOME s)) =>  | 
| 
75123
 
66eb6fdfc244
handle Zipperposition definitions in Isar proof construction
 
blanchet 
parents: 
75122 
diff
changeset
 | 
569  | 
if role = Definition then  | 
| 
 
66eb6fdfc244
handle Zipperposition definitions in Isar proof construction
 
blanchet 
parents: 
75122 
diff
changeset
 | 
570  | 
((num, map fst (find_formula_in_problem phi problem)), phi, role, "", [])  | 
| 57257 | 571  | 
else  | 
| 
75123
 
66eb6fdfc244
handle Zipperposition definitions in Isar proof construction
 
blanchet 
parents: 
75122 
diff
changeset
 | 
572  | 
((num, [s]), phi, role, "", [])  | 
| 
 
66eb6fdfc244
handle Zipperposition definitions in Isar proof construction
 
blanchet 
parents: 
75122 
diff
changeset
 | 
573  | 
| SOME (Inference_Source (rule, deps)) => ((num, []), phi, role, rule, deps)  | 
| 
 
66eb6fdfc244
handle Zipperposition definitions in Isar proof construction
 
blanchet 
parents: 
75122 
diff
changeset
 | 
574  | 
| SOME (Introduced_Source rule) => ((num, []), phi, Definition, rule, [])  | 
| 
 
66eb6fdfc244
handle Zipperposition definitions in Isar proof construction
 
blanchet 
parents: 
75122 
diff
changeset
 | 
575  | 
| SOME Define_Source => ((num, []), phi, Definition, zipperposition_define_rule, [])  | 
| 
 
66eb6fdfc244
handle Zipperposition definitions in Isar proof construction
 
blanchet 
parents: 
75122 
diff
changeset
 | 
576  | 
| _ => ((num, [num]), phi, role, "", []))  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
577  | 
in  | 
| 57257 | 578  | 
[(name, role', phi, rule, map (rpair []) deps)]  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
579  | 
end)  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
580  | 
|
| 75064 | 581  | 
fun parse_tstp_fol_line full problem =  | 
| 77427 | 582  | 
((Scan.this_string tptp_cnf || Scan.this_string tptp_fof || Scan.this_string tptp_tcf  | 
583  | 
      || Scan.this_string tptp_tff) -- $$ "(")
 | 
|
| 50236 | 584  | 
|-- scan_general_id --| $$ "," -- Symbol.scan_ascii_id --| $$ ","  | 
| 75064 | 585  | 
-- (if full then parse_fol_formula || skip_term >> K dummy_phi else skip_term >> K dummy_phi)  | 
586  | 
-- parse_tstp_extra_arguments --| $$ ")" --| $$ "."  | 
|
| 57785 | 587  | 
>> (fn (((num, role0), phi), src) =>  | 
| 42943 | 588  | 
let  | 
| 57785 | 589  | 
val role = role_of_tptp_string role0  | 
590  | 
val ((name, phi), role', rule, deps) =  | 
|
| 42943 | 591  | 
(* Waldmeister isn't exactly helping. *)  | 
| 57785 | 592  | 
(case src of  | 
| 
72967
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
593  | 
SOME (File_Source (_, SOME s)) =>  | 
| 
47927
 
c35238d19bb9
repair the Waldmeister endgame only for Waldmeister proofs
 
blanchet 
parents: 
47926 
diff
changeset
 | 
594  | 
(if s = waldmeister_conjecture_name then  | 
| 54799 | 595  | 
(case find_formula_in_problem (mk_anot phi) problem of  | 
| 57202 | 596  | 
(* Waldmeister hack: Get the original orientation of the equation to avoid  | 
597  | 
confusing Isar. *)  | 
|
| 
47921
 
fc26d5538868
ensure the "show" equation is not reoriented by Waldmeister
 
blanchet 
parents: 
47917 
diff
changeset
 | 
598  | 
[(s, phi')] =>  | 
| 
 
fc26d5538868
ensure the "show" equation is not reoriented by Waldmeister
 
blanchet 
parents: 
47917 
diff
changeset
 | 
599  | 
((num, [s]),  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
600  | 
phi |> not (is_same_formula false [] (mk_anot phi) phi') ? commute_eq)  | 
| 
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
601  | 
| _ => ((num, []), phi))  | 
| 
47921
 
fc26d5538868
ensure the "show" equation is not reoriented by Waldmeister
 
blanchet 
parents: 
47917 
diff
changeset
 | 
602  | 
else  | 
| 
70935
 
535ff1eac86c
removed support for E-ToFoF, which has lost its raison d'etre since E 2.0 supports TF0
 
blanchet 
parents: 
70934 
diff
changeset
 | 
603  | 
((num, [s]), phi),  | 
| 57785 | 604  | 
role, "", [])  | 
| 
72967
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
605  | 
| SOME (File_Source _) =>  | 
| 57785 | 606  | 
(((num, map fst (find_formula_in_problem phi problem)), phi), role, "", [])  | 
| 
72967
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
607  | 
| SOME (Inference_Source (rule, deps)) => (((num, []), phi), role, rule, deps)  | 
| 
75123
 
66eb6fdfc244
handle Zipperposition definitions in Isar proof construction
 
blanchet 
parents: 
75122 
diff
changeset
 | 
608  | 
| SOME (Introduced_Source rule) => (((num, []), phi), Definition, rule, [])  | 
| 
 
66eb6fdfc244
handle Zipperposition definitions in Isar proof construction
 
blanchet 
parents: 
75122 
diff
changeset
 | 
609  | 
| SOME Define_Source => (((num, []), phi), Definition, zipperposition_define_rule, [])  | 
| 
72967
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
610  | 
| _ => (((num, [num]), phi), role, "", []))  | 
| 57785 | 611  | 
|
612  | 
fun mk_step () = (name, role', phi, rule, map (rpair []) deps)  | 
|
| 42943 | 613  | 
in  | 
| 57785 | 614  | 
[(case role' of  | 
| 
56404
 
9cb137ec6ec8
use Z3 TPTP cores rather than proofs since the latter are somewhat broken
 
blanchet 
parents: 
56397 
diff
changeset
 | 
615  | 
Definition =>  | 
| 
 
9cb137ec6ec8
use Z3 TPTP cores rather than proofs since the latter are somewhat broken
 
blanchet 
parents: 
56397 
diff
changeset
 | 
616  | 
(case phi of  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57215 
diff
changeset
 | 
617  | 
                 AAtom (ATerm (("equal", _), _)) =>
 | 
| 
56404
 
9cb137ec6ec8
use Z3 TPTP cores rather than proofs since the latter are somewhat broken
 
blanchet 
parents: 
56397 
diff
changeset
 | 
618  | 
(* Vampire's equality proxy axiom *)  | 
| 
 
9cb137ec6ec8
use Z3 TPTP cores rather than proofs since the latter are somewhat broken
 
blanchet 
parents: 
56397 
diff
changeset
 | 
619  | 
(name, Definition, phi, rule, map (rpair []) deps)  | 
| 57257 | 620  | 
| _ => mk_step ())  | 
| 
56404
 
9cb137ec6ec8
use Z3 TPTP cores rather than proofs since the latter are somewhat broken
 
blanchet 
parents: 
56397 
diff
changeset
 | 
621  | 
| _ => mk_step ())]  | 
| 42943 | 622  | 
end)  | 
| 39452 | 623  | 
|
| 75064 | 624  | 
fun parse_tstp_line full problem =  | 
625  | 
parse_tstp_fol_line full problem  | 
|
626  | 
|| parse_tstp_hol_line full problem  | 
|
| 72398 | 627  | 
|
| 39452 | 628  | 
(**** PARSING OF SPASS OUTPUT ****)  | 
629  | 
||
| 57786 | 630  | 
(* SPASS returns clause references of the form "x.y". We ignore "y". *)  | 
| 39452 | 631  | 
val parse_dot_name = scan_general_id --| $$ "." --| scan_general_id  | 
632  | 
||
633  | 
val parse_spass_annotations =  | 
|
| 54811 | 634  | 
Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name --| Scan.option ($$ ","))) []  | 
| 39452 | 635  | 
|
| 
78696
 
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
 
blanchet 
parents: 
78694 
diff
changeset
 | 
636  | 
(* We ignore the stars and the pluses that follow literals in SPASS's output. *)  | 
| 39602 | 637  | 
fun parse_decorated_atom x =  | 
| 72398 | 638  | 
(parse_fol_atom --| Scan.repeat ($$ "*" || $$ "+" || $$ " ")) x  | 
| 39452 | 639  | 
|
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48130 
diff
changeset
 | 
640  | 
fun mk_horn ([], []) = AAtom (ATerm (("c_False", []), []))
 | 
| 58091 | 641  | 
| mk_horn (neg_lits, pos_lits) = foldr1 (uncurry (mk_aconn AOr)) (map mk_anot neg_lits @ pos_lits)  | 
| 39452 | 642  | 
|
| 39645 | 643  | 
fun parse_horn_clause x =  | 
644  | 
(Scan.repeat parse_decorated_atom --| $$ "|" --| $$ "|"  | 
|
645  | 
-- Scan.repeat parse_decorated_atom --| $$ "-" --| $$ ">"  | 
|
646  | 
-- Scan.repeat parse_decorated_atom  | 
|
647  | 
>> (mk_horn o apfst (op @))) x  | 
|
| 39452 | 648  | 
|
| 46390 | 649  | 
val parse_spass_debug =  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
650  | 
  Scan.option ($$ "(" |-- Scan.repeat (scan_general_id --| Scan.option ($$ ",")) --| $$ ")")
 | 
| 46390 | 651  | 
|
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
652  | 
(* Syntax: <num>[0:<inference><annotations>] <atoms> || <atoms> -> <atoms> .  | 
| 46427 | 653  | 
derived from formulae <ident>* *)  | 
| 
48005
 
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
 
blanchet 
parents: 
47972 
diff
changeset
 | 
654  | 
fun parse_spass_line x =  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
655  | 
(parse_spass_debug |-- scan_general_id --| $$ "[" --| Scan.many1 Symbol.is_digit --| $$ ":"  | 
| 
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
656  | 
-- Symbol.scan_ascii_id -- parse_spass_annotations --| $$ "]" -- parse_horn_clause --| $$ "."  | 
| 
48005
 
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
 
blanchet 
parents: 
47972 
diff
changeset
 | 
657  | 
-- Scan.option (Scan.this_string "derived from formulae "  | 
| 
 
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
 
blanchet 
parents: 
47972 
diff
changeset
 | 
658  | 
|-- Scan.repeat (scan_general_id --| Scan.option ($$ " ")))  | 
| 
 
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
 
blanchet 
parents: 
47972 
diff
changeset
 | 
659  | 
>> (fn ((((num, rule), deps), u), names) =>  | 
| 
56404
 
9cb137ec6ec8
use Z3 TPTP cores rather than proofs since the latter are somewhat broken
 
blanchet 
parents: 
56397 
diff
changeset
 | 
660  | 
[((num, these names), Unknown, u, rule, map (rpair []) deps)])) x  | 
| 45162 | 661  | 
|
| 52077 | 662  | 
fun core_inference inf fact = ((fact, [fact]), Unknown, dummy_phi, inf, [])  | 
663  | 
||
| 75064 | 664  | 
fun parse_line full local_name problem =  | 
| 72398 | 665  | 
if local_name = spassN then parse_spass_line  | 
| 75064 | 666  | 
else parse_tstp_line full problem  | 
| 57154 | 667  | 
|
| 52077 | 668  | 
fun core_of_agsyhol_proof s =  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
669  | 
(case split_lines s of  | 
| 52077 | 670  | 
"The transformed problem consists of the following conjectures:" :: conj ::  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
671  | 
_ :: proof_term :: _ => SOME (unprefix " " conj :: find_enclosed "<<" ">>" proof_term)  | 
| 
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54772 
diff
changeset
 | 
672  | 
| _ => NONE)  | 
| 52077 | 673  | 
|
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42966 
diff
changeset
 | 
674  | 
fun clean_up_dependencies _ [] = []  | 
| 51201 | 675  | 
| clean_up_dependencies seen ((name, role, u, rule, deps) :: steps) =  | 
| 54811 | 676  | 
(name, role, u, rule, map_filter (fn dep => find_first (is_same_atp_step dep) seen) deps) ::  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42966 
diff
changeset
 | 
677  | 
clean_up_dependencies (name :: seen) steps  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42966 
diff
changeset
 | 
678  | 
|
| 42975 | 679  | 
fun clean_up_atp_proof_dependencies proof = clean_up_dependencies [] proof  | 
| 39452 | 680  | 
|
| 50704 | 681  | 
fun map_term_names_in_atp_proof f =  | 
682  | 
let  | 
|
| 54820 | 683  | 
fun map_type (AType ((s, clss), tys)) = AType ((f s, map f clss), map map_type tys)  | 
| 54811 | 684  | 
| map_type (AFun (ty, ty')) = AFun (map_type ty, map_type ty')  | 
685  | 
| map_type (APi (ss, ty)) = APi (map f ss, map_type ty)  | 
|
686  | 
||
687  | 
fun map_term (ATerm ((s, tys), ts)) = ATerm ((f s, map map_type tys), map map_term ts)  | 
|
688  | 
| map_term (AAbs (((s, ty), tm), args)) =  | 
|
689  | 
AAbs (((f s, map_type ty), map_term tm), map map_term args)  | 
|
690  | 
||
| 56397 | 691  | 
fun map_formula (AQuant (q, xs, phi)) = AQuant (q, map (apfst f) xs, map_formula phi)  | 
| 54811 | 692  | 
| map_formula (AConn (c, phis)) = AConn (c, map map_formula phis)  | 
693  | 
| map_formula (AAtom t) = AAtom (map_term t)  | 
|
694  | 
||
| 56397 | 695  | 
fun map_step (name, role, phi, rule, deps) = (name, role, map_formula phi, rule, deps)  | 
| 54811 | 696  | 
in  | 
697  | 
map map_step  | 
|
698  | 
end  | 
|
| 39454 | 699  | 
|
| 57261 | 700  | 
fun nasty_name pool s = Symtab.lookup pool s |> the_default s  | 
| 50704 | 701  | 
|
| 39454 | 702  | 
fun nasty_atp_proof pool =  | 
| 50704 | 703  | 
not (Symtab.is_empty pool) ? map_term_names_in_atp_proof (nasty_name pool)  | 
| 39454 | 704  | 
|
| 72355 | 705  | 
fun string_of_list f xs = enclose "[" "]" (commas (map f xs))  | 
706  | 
||
707  | 
fun string_of_atp_step_name (s, ss) = "(" ^ s ^ ", " ^ string_of_list I ss ^ ")"
 | 
|
708  | 
||
709  | 
fun string_of_atp_step f g (name, role, x, y, names) =  | 
|
710  | 
let  | 
|
711  | 
val name' = string_of_atp_step_name name  | 
|
712  | 
val role' = ATP_Problem.tptp_string_of_role role  | 
|
713  | 
val x' = f x  | 
|
714  | 
val y' = g y  | 
|
715  | 
val names' = string_of_list string_of_atp_step_name names  | 
|
716  | 
in  | 
|
717  | 
    "(" ^ name' ^ ", " ^ role' ^ ", " ^ x' ^ ", " ^ y' ^ ", " ^ names' ^ ")"
 | 
|
718  | 
end  | 
|
719  | 
||
| 75064 | 720  | 
fun parse_proof full local_name problem =  | 
| 
72399
 
f8900a5ad4a7
drop obsolete ad hoc support for Satallax isar proof reconstruction
 
desharna 
parents: 
72398 
diff
changeset
 | 
721  | 
strip_spaces_except_between_idents  | 
| 
 
f8900a5ad4a7
drop obsolete ad hoc support for Satallax isar proof reconstruction
 
desharna 
parents: 
72398 
diff
changeset
 | 
722  | 
#> raw_explode  | 
| 
 
f8900a5ad4a7
drop obsolete ad hoc support for Satallax isar proof reconstruction
 
desharna 
parents: 
72398 
diff
changeset
 | 
723  | 
#> Scan.error (!! (fn _ => raise UNRECOGNIZED_ATP_PROOF ())  | 
| 75064 | 724  | 
(Scan.finite Symbol.stopper (Scan.repeats1 (parse_line full local_name problem))))  | 
725  | 
#> (fn (proof, ss) => if null ss then proof else raise UNRECOGNIZED_ATP_PROOF ())  | 
|
| 
72399
 
f8900a5ad4a7
drop obsolete ad hoc support for Satallax isar proof reconstruction
 
desharna 
parents: 
72398 
diff
changeset
 | 
726  | 
|
| 75064 | 727  | 
fun atp_proof_of_tstplike_proof _ _ _ "" = []  | 
728  | 
| atp_proof_of_tstplike_proof full local_prover problem tstp =  | 
|
| 
72399
 
f8900a5ad4a7
drop obsolete ad hoc support for Satallax isar proof reconstruction
 
desharna 
parents: 
72398 
diff
changeset
 | 
729  | 
(case core_of_agsyhol_proof tstp of  | 
| 
 
f8900a5ad4a7
drop obsolete ad hoc support for Satallax isar proof reconstruction
 
desharna 
parents: 
72398 
diff
changeset
 | 
730  | 
SOME facts => facts |> map (core_inference agsyhol_core_rule)  | 
| 
 
f8900a5ad4a7
drop obsolete ad hoc support for Satallax isar proof reconstruction
 
desharna 
parents: 
72398 
diff
changeset
 | 
731  | 
| NONE =>  | 
| 75064 | 732  | 
tstp  | 
733  | 
|> parse_proof full local_prover problem  | 
|
| 
72967
 
11de287ed481
tweaked tptp parsing when source info is missing
 
desharna 
parents: 
72689 
diff
changeset
 | 
734  | 
|> local_prover = vampireN ? perhaps (try (sort (vampire_step_name_ord o apply2 #1))))  | 
| 
72399
 
f8900a5ad4a7
drop obsolete ad hoc support for Satallax isar proof reconstruction
 
desharna 
parents: 
72398 
diff
changeset
 | 
735  | 
|
| 
77418
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
736  | 
val e_symbol_prefixes = ["esk", "epred"]  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
737  | 
|
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
738  | 
fun exists_name_in_term pred =  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
739  | 
let  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
740  | 
fun ex_name (ATerm ((s, _), tms)) = pred s orelse exists ex_name tms  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
741  | 
| ex_name (AAbs ((_, tm), tms)) = exists ex_name (tm :: tms)  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
742  | 
in ex_name end  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
743  | 
|
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
744  | 
fun exists_name_in_formula pred phi =  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
745  | 
formula_fold NONE (fn _ => fn tm => fn ex => ex orelse exists_name_in_term pred tm) phi false  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
746  | 
|
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
747  | 
fun exists_symbol_in_formula prefixes =  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
748  | 
exists_name_in_formula (fn s => exists (fn prefix => String.isPrefix prefix s) prefixes)  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
749  | 
|
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
750  | 
fun atp_abduce_candidates_of_output local_prover problem output =  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
751  | 
let  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
752  | 
(* Truncate too large output to avoid memory issues. *)  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
753  | 
val max_size = 1000000  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
754  | 
val output =  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
755  | 
if String.size output > max_size then  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
756  | 
String.substring (output, 0, max_size)  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
757  | 
else  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
758  | 
output  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
759  | 
|
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
760  | 
fun fold_extract accum [] = accum  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
761  | 
| fold_extract accum (line :: lines) =  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
762  | 
if String.isSubstring "# info" line  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
763  | 
andalso String.isSubstring "negated_conjecture" line then  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
764  | 
if String.isSubstring ", 0, 0," line then  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
765  | 
(* This pattern occurs in the "info()" comment of an E clause that directly emerges from  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
766  | 
the conjecture. We don't want to tell the user that they can prove "P" by assuming  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
767  | 
"P". *)  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
768  | 
fold_extract accum lines  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
769  | 
else  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
770  | 
let  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
771  | 
val clean_line =  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
772  | 
(case space_explode "#" line of  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
773  | 
[] => ""  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
774  | 
| before_hash :: _ => before_hash)  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
775  | 
in  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
776  | 
(case try (parse_proof true local_prover problem) clean_line of  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
777  | 
SOME [(_, _, phi, _, _)] =>  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
778  | 
if local_prover = eN andalso exists_symbol_in_formula e_symbol_prefixes phi then  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
779  | 
fold_extract accum lines  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
780  | 
else  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
781  | 
fold_extract (phi :: accum) lines  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
782  | 
| _ => fold_extract accum lines)  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
783  | 
end  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
784  | 
else  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
785  | 
fold_extract accum lines  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
786  | 
in  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
787  | 
fold_extract [] (split_lines output)  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
788  | 
end  | 
| 
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77272 
diff
changeset
 | 
789  | 
|
| 39452 | 790  | 
end;  |