author | fleury |
Mon, 16 Jun 2014 16:18:15 +0200 | |
changeset 57255 | 488046fdda59 |
parent 57199 | 472360558b22 |
child 57256 | cf43583f9bb9 |
permissions | -rw-r--r-- |
46320 | 1 |
(* Title: HOL/Tools/ATP/atp_proof_reconstruct.ML |
38027 | 2 |
Author: Lawrence C. Paulson, Cambridge University Computer Laboratory |
3 |
Author: Claire Quigley, Cambridge University Computer Laboratory |
|
36392 | 4 |
Author: Jasmin Blanchette, TU Muenchen |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
5 |
|
49914 | 6 |
Basic proof reconstruction from ATP proofs. |
33310 | 7 |
*) |
8 |
||
46320 | 9 |
signature ATP_PROOF_RECONSTRUCT = |
24425
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
10 |
sig |
54811 | 11 |
type 'a atp_type = 'a ATP_Problem.atp_type |
53586
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
52758
diff
changeset
|
12 |
type ('a, 'b) atp_term = ('a, 'b) ATP_Problem.atp_term |
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
52758
diff
changeset
|
13 |
type ('a, 'b, 'c, 'd) atp_formula = ('a, 'b, 'c, 'd) ATP_Problem.atp_formula |
54500 | 14 |
type stature = ATP_Problem_Generate.stature |
54772 | 15 |
type atp_step_name = ATP_Proof.atp_step_name |
54499 | 16 |
type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step |
17 |
type 'a atp_proof = 'a ATP_Proof.atp_proof |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
18 |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
19 |
val metisN : string |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
20 |
val full_typesN : string |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
21 |
val partial_typesN : string |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
22 |
val no_typesN : string |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
23 |
val really_full_type_enc : string |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
24 |
val full_type_enc : string |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
25 |
val partial_type_enc : string |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
26 |
val no_type_enc : string |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
27 |
val full_type_encs : string list |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
28 |
val partial_type_encs : string list |
54500 | 29 |
val default_metis_lam_trans : string |
54772 | 30 |
|
49983
33e18e9916a8
use metaquantification when possible in Isar proofs
blanchet
parents:
49914
diff
changeset
|
31 |
val forall_of : term -> term -> term |
33e18e9916a8
use metaquantification when possible in Isar proofs
blanchet
parents:
49914
diff
changeset
|
32 |
val exists_of : term -> term -> term |
55285 | 33 |
val type_enc_aliases : (string * string list) list |
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
34 |
val unalias_type_enc : string -> string list |
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
35 |
val term_of_atp : Proof.context -> ATP_Problem.atp_format -> ATP_Problem_Generate.type_enc -> |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
36 |
bool -> int Symtab.table -> typ option -> (string, string atp_type) atp_term -> term |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
37 |
val prop_of_atp : Proof.context -> ATP_Problem.atp_format -> ATP_Problem_Generate.type_enc -> |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
38 |
bool -> int Symtab.table -> |
54811 | 39 |
(string, string, (string, string atp_type) atp_term, string) atp_formula -> term |
54499 | 40 |
|
54500 | 41 |
val used_facts_in_atp_proof : |
54507 | 42 |
Proof.context -> (string * stature) list vector -> string atp_proof -> (string * stature) list |
43 |
val used_facts_in_unsound_atp_proof : Proof.context -> (string * stature) list vector -> |
|
44 |
'a atp_proof -> string list option |
|
55257 | 45 |
val atp_proof_prefers_lifting : string atp_proof -> bool |
54500 | 46 |
val is_typed_helper_used_in_atp_proof : string atp_proof -> bool |
54772 | 47 |
val replace_dependencies_in_line : atp_step_name * atp_step_name list -> ('a, 'b) atp_step -> |
48 |
('a, 'b) atp_step |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
49 |
val termify_atp_proof : Proof.context -> string -> ATP_Problem.atp_format -> |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
50 |
ATP_Problem_Generate.type_enc -> string Symtab.table -> (string * term) list -> |
54499 | 51 |
int Symtab.table -> string atp_proof -> (term, string) atp_step list |
54772 | 52 |
val introduce_spass_skolem : (term, string) atp_step list -> (term, string) atp_step list |
54505 | 53 |
val factify_atp_proof : (string * 'a) list vector -> term list -> term -> |
54 |
(term, string) atp_step list -> (term, string) atp_step list |
|
24425
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
55 |
end; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
56 |
|
46320 | 57 |
structure ATP_Proof_Reconstruct : ATP_PROOF_RECONSTRUCT = |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
58 |
struct |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
59 |
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43062
diff
changeset
|
60 |
open ATP_Util |
38028 | 61 |
open ATP_Problem |
39452 | 62 |
open ATP_Proof |
46320 | 63 |
open ATP_Problem_Generate |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
64 |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
65 |
val metisN = "metis" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
66 |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
67 |
val full_typesN = "full_types" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
68 |
val partial_typesN = "partial_types" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
69 |
val no_typesN = "no_types" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
70 |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
71 |
val really_full_type_enc = "mono_tags" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
72 |
val full_type_enc = "poly_guards_query" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
73 |
val partial_type_enc = "poly_args" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
74 |
val no_type_enc = "erased" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
75 |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
76 |
val full_type_encs = [full_type_enc, really_full_type_enc] |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
77 |
val partial_type_encs = partial_type_enc :: full_type_encs |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
78 |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
79 |
val type_enc_aliases = |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
80 |
[(full_typesN, full_type_encs), |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
81 |
(partial_typesN, partial_type_encs), |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
82 |
(no_typesN, [no_type_enc])] |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
83 |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
84 |
fun unalias_type_enc s = |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
85 |
AList.lookup (op =) type_enc_aliases s |> the_default [s] |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
86 |
|
54500 | 87 |
val default_metis_lam_trans = combsN |
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
88 |
|
50670
eaa540986291
properly take the existential closure of skolems
blanchet
parents:
49983
diff
changeset
|
89 |
fun term_name' (Var ((s, _), _)) = perhaps (try Name.dest_skolem) s |
52034
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
90 |
| term_name' _ = "" |
50670
eaa540986291
properly take the existential closure of skolems
blanchet
parents:
49983
diff
changeset
|
91 |
|
eaa540986291
properly take the existential closure of skolems
blanchet
parents:
49983
diff
changeset
|
92 |
fun lambda' v = Term.lambda_name (term_name' v, v) |
eaa540986291
properly take the existential closure of skolems
blanchet
parents:
49983
diff
changeset
|
93 |
|
eaa540986291
properly take the existential closure of skolems
blanchet
parents:
49983
diff
changeset
|
94 |
fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda' v t |
eaa540986291
properly take the existential closure of skolems
blanchet
parents:
49983
diff
changeset
|
95 |
fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda' v t |
39425 | 96 |
|
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
97 |
fun make_tfree ctxt w = |
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
98 |
let val ww = "'" ^ w in |
56254 | 99 |
TFree (ww, the_default @{sort type} (Variable.def_sort ctxt (ww, ~1))) |
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
100 |
end |
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
101 |
|
54820 | 102 |
exception ATP_CLASS of string list |
54818 | 103 |
exception ATP_TYPE of string atp_type list |
54811 | 104 |
exception ATP_TERM of (string, string atp_type) atp_term list |
53586
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
52758
diff
changeset
|
105 |
exception ATP_FORMULA of |
54811 | 106 |
(string, string, (string, string atp_type) atp_term, string) atp_formula list |
37991 | 107 |
exception SAME of unit |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
108 |
|
54820 | 109 |
fun class_of_atp_class cls = |
110 |
(case unprefix_and_unascii class_prefix cls of |
|
111 |
SOME s => s |
|
112 |
| NONE => raise ATP_CLASS [cls]) |
|
113 |
||
54811 | 114 |
(* Type variables are given the basic sort "HOL.type". Some will later be constrained by information |
115 |
from type literals, or by type inference. *) |
|
54820 | 116 |
fun typ_of_atp_type ctxt (ty as AType ((a, clss), tys)) = |
54818 | 117 |
let val Ts = map (typ_of_atp_type ctxt) tys in |
54789 | 118 |
(case unprefix_and_unascii type_const_prefix a of |
37991 | 119 |
SOME b => Type (invert_const b, Ts) |
120 |
| NONE => |
|
54818 | 121 |
if not (null tys) then |
122 |
raise ATP_TYPE [ty] (* only "tconst"s have type arguments *) |
|
54789 | 123 |
else |
124 |
(case unprefix_and_unascii tfree_prefix a of |
|
125 |
SOME b => make_tfree ctxt b |
|
126 |
| NONE => |
|
127 |
(* The term could be an Isabelle variable or a variable from the ATP, say "X1" or "_5018". |
|
128 |
Sometimes variables from the ATP are indistinguishable from Isabelle variables, which |
|
129 |
forces us to use a type parameter in all cases. *) |
|
54821 | 130 |
Type_Infer.param 0 ("'" ^ perhaps (unprefix_and_unascii tvar_prefix) a, |
56254 | 131 |
(if null clss then @{sort type} else map class_of_atp_class clss)))) |
37991 | 132 |
end |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
133 |
|
54820 | 134 |
fun atp_type_of_atp_term (ATerm ((s, _), us)) = AType ((s, []), map atp_type_of_atp_term us) |
54818 | 135 |
|
136 |
fun typ_of_atp_term ctxt = typ_of_atp_type ctxt o atp_type_of_atp_term |
|
137 |
||
54789 | 138 |
(* Type class literal applied to a type. Returns triple of polarity, class, type. *) |
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51878
diff
changeset
|
139 |
fun type_constraint_of_term ctxt (u as ATerm ((a, _), us)) = |
54818 | 140 |
(case (unprefix_and_unascii class_prefix a, map (typ_of_atp_term ctxt) us) of |
42606
0c76cf483899
show sorts not just types in Isar proofs + tuning
blanchet
parents:
42595
diff
changeset
|
141 |
(SOME b, [T]) => (b, T) |
54789 | 142 |
| _ => raise ATP_TERM [u]) |
38014 | 143 |
|
43907
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
144 |
(* Accumulate type constraints in a formula: negative type literals. *) |
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
145 |
fun add_var (key, z) = Vartab.map_default (key, []) (cons z) |
42606
0c76cf483899
show sorts not just types in Isar proofs + tuning
blanchet
parents:
42595
diff
changeset
|
146 |
fun add_type_constraint false (cl, TFree (a ,_)) = add_var ((a, ~1), cl) |
0c76cf483899
show sorts not just types in Isar proofs + tuning
blanchet
parents:
42595
diff
changeset
|
147 |
| add_type_constraint false (cl, TVar (ix, _)) = add_var (ix, cl) |
0c76cf483899
show sorts not just types in Isar proofs + tuning
blanchet
parents:
42595
diff
changeset
|
148 |
| add_type_constraint _ _ = I |
38014 | 149 |
|
55185
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
150 |
fun repair_var_name_raw s = |
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
151 |
let |
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
152 |
fun subscript_name s n = s ^ nat_subscript n |
50704 | 153 |
val s = s |> String.map Char.toLower |
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
154 |
in |
54789 | 155 |
(case space_explode "_" s of |
156 |
[_] => |
|
157 |
(case take_suffix Char.isDigit (String.explode s) of |
|
158 |
(cs1 as _ :: _, cs2 as _ :: _) => |
|
159 |
subscript_name (String.implode cs1) (the (Int.fromString (String.implode cs2))) |
|
160 |
| (_, _) => s) |
|
161 |
| [s1, s2] => (case Int.fromString s2 of SOME n => subscript_name s1 n | NONE => s) |
|
162 |
| _ => s) |
|
36486
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents:
36485
diff
changeset
|
163 |
end |
43182
649bada59658
slacker version of "fastype_of", in case a function has dummy type
blanchet
parents:
43176
diff
changeset
|
164 |
|
55185
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
165 |
fun repair_var_name textual s = |
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
166 |
(case unprefix_and_unascii schematic_var_prefix s of |
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
167 |
SOME s => s |
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
168 |
| NONE => s |> textual ? repair_var_name_raw) |
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
169 |
|
54789 | 170 |
(* The number of type arguments of a constant, zero if it's monomorphic. For (instances of) Skolem |
171 |
pseudoconstants, this information is encoded in the constant name. *) |
|
52125
ac7830871177
improved handling of free variables' types in Isar proofs
blanchet
parents:
52034
diff
changeset
|
172 |
fun robust_const_num_type_args thy s = |
43907
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
173 |
if String.isPrefix skolem_const_prefix s then |
46711
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
wenzelm
parents:
46435
diff
changeset
|
174 |
s |> Long_Name.explode |> List.last |> Int.fromString |> the |
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
175 |
else if String.isPrefix lam_lifted_prefix s then |
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
176 |
if String.isPrefix lam_lifted_poly_prefix s then 2 else 0 |
43907
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
177 |
else |
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
178 |
(s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length |
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
179 |
|
56254 | 180 |
fun slack_fastype_of t = fastype_of t handle TERM _ => Type_Infer.anyT @{sort type} |
43182
649bada59658
slacker version of "fastype_of", in case a function has dummy type
blanchet
parents:
43176
diff
changeset
|
181 |
|
51192
4dc6bb65c3c3
slacker comparison for Skolems, to avoid trivial equations
blanchet
parents:
50704
diff
changeset
|
182 |
(* Cope with "tt(X) = X" atoms, where "X" is existentially quantified. *) |
4dc6bb65c3c3
slacker comparison for Skolems, to avoid trivial equations
blanchet
parents:
50704
diff
changeset
|
183 |
fun loose_aconv (Free (s, _), Free (s', _)) = s = s' |
4dc6bb65c3c3
slacker comparison for Skolems, to avoid trivial equations
blanchet
parents:
50704
diff
changeset
|
184 |
| loose_aconv (t, t') = t aconv t' |
4dc6bb65c3c3
slacker comparison for Skolems, to avoid trivial equations
blanchet
parents:
50704
diff
changeset
|
185 |
|
54772 | 186 |
val spass_skolem_prefix = "sk" (* "skc" or "skf" *) |
187 |
val vampire_skolem_prefix = "sK" |
|
50703
76a2e506c125
swap Vampire's Skolem arguments to bring them in line with what E and metis's new skolemizer do (helps Isar proof reconstruction in some cases)
blanchet
parents:
50693
diff
changeset
|
188 |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
189 |
|
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
190 |
(* Higher-order translation. Variables are typed (although we don't use that information). Lambdas |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
191 |
are typed. *) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
192 |
fun term_of_atp_ho ctxt textual sym_tab = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
193 |
let |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
194 |
val thy = Proof_Context.theory_of ctxt |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
195 |
fun norm_var_types var T' t = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
196 |
let |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
197 |
fun norm_term (ATerm ((x, ty), l)) = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
198 |
ATerm((x, if x = var then [T'] else ty), List.map norm_term l) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
199 |
| norm_term (AAbs(((x, ty), term), l)) = AAbs(((x, ty), term), List.map norm_term l) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
200 |
in |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
201 |
norm_term t |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
202 |
end |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
203 |
fun mk_op_of_tptp_operator s = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
204 |
if s = tptp_or then HOLogic.mk_disj |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
205 |
else if s = tptp_and then HOLogic.mk_conj |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
206 |
else if s = tptp_implies then HOLogic.mk_imp |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
207 |
else if s = tptp_iff orelse s = tptp_equal then HOLogic.mk_eq |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
208 |
else if s = tptp_not_iff orelse s = tptp_not_equal then HOLogic.mk_eq #> HOLogic.mk_not |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
209 |
else if s = tptp_if then HOLogic.mk_imp #> HOLogic.mk_not |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
210 |
else if s = tptp_not_and then HOLogic.mk_conj #> HOLogic.mk_not |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
211 |
else if s = tptp_not_or then HOLogic.mk_disj #> HOLogic.mk_not |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
212 |
else raise Fail ("unknown operator: " ^ s) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
213 |
fun mk_quant_of_tptp_quant s = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
214 |
if s = tptp_ho_exists then HOLogic.mk_exists |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
215 |
else if s = tptp_ho_forall then HOLogic.mk_all |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
216 |
else raise Fail ("unknown quantifier: " ^ s) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
217 |
val var_index = if textual then 0 else 1 |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
218 |
|
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
219 |
fun do_term opt_T u = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
220 |
(case u of |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
221 |
AAbs(((var, ty), term), []) => |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
222 |
let val typ = typ_of_atp_type ctxt ty in |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
223 |
absfree (repair_var_name textual var, typ) (do_term NONE (norm_var_types var ty term)) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
224 |
end |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
225 |
| ATerm ((s, tys), us) => |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
226 |
if s = "" |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
227 |
then error "Isar proof reconstruction failed because the ATP proof \ |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
228 |
\contains unparsable material." |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
229 |
else if String.isPrefix native_type_prefix s then |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
230 |
@{const True} (* ignore TPTP type information *) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
231 |
else if s = tptp_equal then |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
232 |
let val ts = map (do_term NONE) us in |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
233 |
if textual andalso length ts = 2 andalso loose_aconv (hd ts, List.last ts) then |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
234 |
@{const True} |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
235 |
else |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
236 |
list_comb (Const (@{const_name HOL.eq}, Type_Infer.anyT @{sort type}), ts) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
237 |
end |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
238 |
else if s = tptp_app then |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
239 |
let val ts = map (do_term NONE) us in |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
240 |
hd ts $ List.last ts |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
241 |
end |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
242 |
else if s = tptp_not then |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
243 |
let val ts = map (do_term NONE) us in |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
244 |
List.last ts |> HOLogic.mk_not |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
245 |
end |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
246 |
else if s = tptp_ho_forall orelse s = tptp_ho_exists then |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
247 |
(case us of |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
248 |
[AAbs (((var, ty), term), [])] => |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
249 |
let val typ = typ_of_atp_type ctxt ty in |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
250 |
(repair_var_name textual var, typ , do_term NONE (norm_var_types var ty term)) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
251 |
|> mk_quant_of_tptp_quant s |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
252 |
end |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
253 |
| [] => if s = tptp_ho_exists then HOLogic.exists_const dummyT |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
254 |
else HOLogic.all_const dummyT) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
255 |
else if List.exists (curry (op=) s) tptp_binary_op_list then |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
256 |
let val ts = map (do_term NONE) us in |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
257 |
(mk_op_of_tptp_operator s) (hd ts, List.last ts) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
258 |
end |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
259 |
else |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
260 |
if us <> [] then |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
261 |
let |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
262 |
val applicant_list = List.map (do_term NONE) us |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
263 |
val opt_typ = SOME (fold_rev |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
264 |
(fn t1 => fn t2 => slack_fastype_of t1 --> t2) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
265 |
applicant_list (the_default dummyT opt_T)) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
266 |
val func_name = do_term opt_typ (ATerm ((s, tys), [])) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
267 |
in |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
268 |
foldl1 (op$) (func_name :: applicant_list) end |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
269 |
else (*FIXME: clean (remove stuff related to vampire and other provers)*) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
270 |
(case unprefix_and_unascii const_prefix s of |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
271 |
SOME s' => |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
272 |
let |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
273 |
val ((s', s''), mangled_us) = s' |> unmangled_const |>> `invert_const |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
274 |
val new_skolem = String.isPrefix new_skolem_const_prefix s'' |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
275 |
val num_ty_args = length us - the_default 0 (Symtab.lookup sym_tab s) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
276 |
val (type_us, term_us) = chop num_ty_args us |>> append mangled_us |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
277 |
val term_ts = map (do_term NONE) term_us |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
278 |
val Ts = map (typ_of_atp_type ctxt) tys @ map (typ_of_atp_term ctxt) type_us |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
279 |
val T = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
280 |
(if not (null Ts) andalso robust_const_num_type_args thy s' = length Ts then |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
281 |
if new_skolem then |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
282 |
SOME (Type_Infer.paramify_vars (tl Ts ---> hd Ts)) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
283 |
else if textual then |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
284 |
try (Sign.const_instance thy) (s', Ts) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
285 |
else |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
286 |
NONE |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
287 |
else |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
288 |
NONE) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
289 |
|> (fn SOME T => T |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
290 |
| NONE => |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
291 |
map slack_fastype_of term_ts ---> |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
292 |
the_default (Type_Infer.anyT @{sort type}) opt_T) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
293 |
val t = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
294 |
if new_skolem then Var ((new_skolem_var_name_of_const s'', var_index), T) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
295 |
else Const (unproxify_const s', T) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
296 |
in list_comb (t, term_ts) end |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
297 |
| NONE => (* a free or schematic variable *) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
298 |
let |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
299 |
(* This assumes that distinct names are mapped to distinct names by |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
300 |
"Variable.variant_frees". This does not hold in general but should hold for |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
301 |
ATP-generated Skolem function names, since these end with a digit and |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
302 |
"variant_frees" appends letters. *) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
303 |
fun fresh_up s = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
304 |
[(s, ())] |> Variable.variant_frees ctxt [] |> hd |> fst |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
305 |
val ts = map (do_term NONE) us |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
306 |
val T = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
307 |
(case opt_T of |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
308 |
SOME T => map slack_fastype_of ts ---> T |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
309 |
| NONE => |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
310 |
map slack_fastype_of ts ---> |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
311 |
(case tys of [ty] => typ_of_atp_type ctxt ty | _ => Type_Infer.anyT |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
312 |
@{sort type})) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
313 |
val t = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
314 |
(case unprefix_and_unascii fixed_var_prefix s of |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
315 |
SOME s => Free (s, T) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
316 |
| NONE => |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
317 |
if textual andalso not (is_tptp_variable s) then |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
318 |
Free (s |> textual ? (repair_var_name_raw #> fresh_up), T) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
319 |
else |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
320 |
Var ((repair_var_name textual s, var_index), T)) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
321 |
in list_comb (t, ts) end)) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
322 |
in do_term end |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
323 |
|
56254 | 324 |
(* First-order translation. No types are known for variables. "Type_Infer.anyT @{sort type}" |
325 |
should allow them to be inferred. *) |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
326 |
fun term_of_atp_fo ctxt textual sym_tab = |
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
327 |
let |
43135
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents:
43131
diff
changeset
|
328 |
val thy = Proof_Context.theory_of ctxt |
54789 | 329 |
(* For Metis, we use 1 rather than 0 because variable references in clauses may otherwise |
330 |
conflict with variable constraints in the goal. At least, type inference often fails |
|
331 |
otherwise. See also "axiom_inference" in "Metis_Reconstruct". *) |
|
43094 | 332 |
val var_index = if textual then 0 else 1 |
57199 | 333 |
|
43131
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents:
43130
diff
changeset
|
334 |
fun do_term extra_ts opt_T u = |
54789 | 335 |
(case u of |
54818 | 336 |
ATerm ((s, tys), us) => |
57199 | 337 |
if s = "" then |
338 |
error "Isar proof reconstruction failed because the ATP proof contains unparsable \ |
|
339 |
\material." |
|
51878 | 340 |
else if String.isPrefix native_type_prefix s then |
42962
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
blanchet
parents:
42943
diff
changeset
|
341 |
@{const True} (* ignore TPTP type information *) |
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44399
diff
changeset
|
342 |
else if s = tptp_equal then |
43093 | 343 |
let val ts = map (do_term [] NONE) us in |
54822 | 344 |
if textual andalso length ts = 2 andalso loose_aconv (hd ts, List.last ts) then |
39106 | 345 |
@{const True} |
346 |
else |
|
56254 | 347 |
list_comb (Const (@{const_name HOL.eq}, Type_Infer.anyT @{sort type}), ts) |
39106 | 348 |
end |
54789 | 349 |
else |
350 |
(case unprefix_and_unascii const_prefix s of |
|
351 |
SOME s' => |
|
54818 | 352 |
let val ((s', s''), mangled_us) = s' |> unmangled_const |>> `invert_const in |
54789 | 353 |
if s' = type_tag_name then |
354 |
(case mangled_us @ us of |
|
54818 | 355 |
[typ_u, term_u] => do_term extra_ts (SOME (typ_of_atp_term ctxt typ_u)) term_u |
54789 | 356 |
| _ => raise ATP_TERM us) |
357 |
else if s' = predicator_name then |
|
358 |
do_term [] (SOME @{typ bool}) (hd us) |
|
359 |
else if s' = app_op_name then |
|
360 |
let val extra_t = do_term [] NONE (List.last us) in |
|
361 |
do_term (extra_t :: extra_ts) |
|
57199 | 362 |
(case opt_T of SOME T => SOME (slack_fastype_of extra_t --> T) | NONE => NONE) |
363 |
(nth us (length us - 2)) |
|
54789 | 364 |
end |
365 |
else if s' = type_guard_name then |
|
366 |
@{const True} (* ignore type predicates *) |
|
367 |
else |
|
368 |
let |
|
369 |
val new_skolem = String.isPrefix new_skolem_const_prefix s'' |
|
54818 | 370 |
val num_ty_args = length us - the_default 0 (Symtab.lookup sym_tab s) |
371 |
val (type_us, term_us) = chop num_ty_args us |>> append mangled_us |
|
54789 | 372 |
val term_ts = map (do_term [] NONE) term_us |
54818 | 373 |
|
374 |
val Ts = map (typ_of_atp_type ctxt) tys @ map (typ_of_atp_term ctxt) type_us |
|
54789 | 375 |
val T = |
54818 | 376 |
(if not (null Ts) andalso robust_const_num_type_args thy s' = length Ts then |
57199 | 377 |
if new_skolem then SOME (Type_Infer.paramify_vars (tl Ts ---> hd Ts)) |
378 |
else if textual then try (Sign.const_instance thy) (s', Ts) |
|
379 |
else NONE |
|
54789 | 380 |
else |
381 |
NONE) |
|
382 |
|> (fn SOME T => T |
|
383 |
| NONE => |
|
56254 | 384 |
map slack_fastype_of term_ts ---> |
57199 | 385 |
the_default (Type_Infer.anyT @{sort type}) opt_T) |
54789 | 386 |
val t = |
387 |
if new_skolem then Var ((new_skolem_var_name_of_const s'', var_index), T) |
|
388 |
else Const (unproxify_const s', T) |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
389 |
in |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
390 |
list_comb (t, term_ts @ extra_ts) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
391 |
end |
54789 | 392 |
end |
393 |
| NONE => (* a free or schematic variable *) |
|
394 |
let |
|
395 |
(* This assumes that distinct names are mapped to distinct names by |
|
54822 | 396 |
"Variable.variant_frees". This does not hold in general but should hold for |
397 |
ATP-generated Skolem function names, since these end with a digit and |
|
398 |
"variant_frees" appends letters. *) |
|
57199 | 399 |
fun fresh_up s = [(s, ())] |> Variable.variant_frees ctxt [] |> hd |> fst |
400 |
||
54789 | 401 |
val term_ts = |
402 |
map (do_term [] NONE) us |
|
403 |
(* SPASS (3.8ds) and Vampire (2.6) pass arguments to Skolem functions in reverse |
|
404 |
order, which is incompatible with the new Metis skolemizer. *) |
|
405 |
|> exists (fn pre => String.isPrefix pre s) |
|
406 |
[spass_skolem_prefix, vampire_skolem_prefix] ? rev |
|
407 |
val ts = term_ts @ extra_ts |
|
408 |
val T = |
|
409 |
(case opt_T of |
|
410 |
SOME T => map slack_fastype_of term_ts ---> T |
|
54822 | 411 |
| NONE => |
412 |
map slack_fastype_of ts ---> |
|
56254 | 413 |
(case tys of [ty] => typ_of_atp_type ctxt ty | _ => Type_Infer.anyT @{sort type})) |
54789 | 414 |
val t = |
415 |
(case unprefix_and_unascii fixed_var_prefix s of |
|
416 |
SOME s => Free (s, T) |
|
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
417 |
| NONE => |
55185
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
418 |
if textual andalso not (is_tptp_variable s) then |
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
419 |
Free (s |> textual ? (repair_var_name_raw #> fresh_up), T) |
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
420 |
else |
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
421 |
Var ((repair_var_name textual s, var_index), T)) |
54789 | 422 |
in list_comb (t, ts) end)) |
43093 | 423 |
in do_term [] end |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
424 |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
425 |
fun term_of_atp ctxt (ATP_Problem.THF _) type_enc = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
426 |
if ATP_Problem_Generate.is_type_enc_higher_order type_enc |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
427 |
then term_of_atp_ho ctxt |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
428 |
else error "Unsupported Isar reconstruction." |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
429 |
| term_of_atp ctxt _ type_enc = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
430 |
if not (ATP_Problem_Generate.is_type_enc_higher_order type_enc) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
431 |
then term_of_atp_fo ctxt |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
432 |
else error "Unsupported Isar reconstruction." |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
433 |
|
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
434 |
fun term_of_atom ctxt format type_enc textual sym_tab pos (u as ATerm ((s, _), _)) = |
38014 | 435 |
if String.isPrefix class_prefix s then |
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51878
diff
changeset
|
436 |
add_type_constraint pos (type_constraint_of_term ctxt u) |
38014 | 437 |
#> pair @{const True} |
438 |
else |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
439 |
pair (term_of_atp ctxt format type_enc textual sym_tab (SOME @{typ bool}) u) |
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
440 |
|
37991 | 441 |
(* Update schematic type variables with detected sort constraints. It's not |
42563 | 442 |
totally clear whether this code is necessary. *) |
38014 | 443 |
fun repair_tvar_sorts (t, tvar_tab) = |
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
444 |
let |
37991 | 445 |
fun do_type (Type (a, Ts)) = Type (a, map do_type Ts) |
446 |
| do_type (TVar (xi, s)) = |
|
447 |
TVar (xi, the_default s (Vartab.lookup tvar_tab xi)) |
|
448 |
| do_type (TFree z) = TFree z |
|
449 |
fun do_term (Const (a, T)) = Const (a, do_type T) |
|
450 |
| do_term (Free (a, T)) = Free (a, do_type T) |
|
451 |
| do_term (Var (xi, T)) = Var (xi, do_type T) |
|
452 |
| do_term (t as Bound _) = t |
|
453 |
| do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t) |
|
454 |
| do_term (t1 $ t2) = do_term t1 $ do_term t2 |
|
455 |
in t |> not (Vartab.is_empty tvar_tab) ? do_term end |
|
456 |
||
39425 | 457 |
fun quantify_over_var quant_of var_s t = |
458 |
let |
|
55185
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
459 |
val vars = filter (fn ((s, _), _) => s = var_s) (Term.add_vars t []) |
52758
7ffcd6f2890d
avoid duplicating Var when the types do not quite fit -- since this step occurs before type inference
blanchet
parents:
52125
diff
changeset
|
460 |
val normTs = vars |> AList.group (op =) |> map (apsnd hd) |
7ffcd6f2890d
avoid duplicating Var when the types do not quite fit -- since this step occurs before type inference
blanchet
parents:
52125
diff
changeset
|
461 |
fun norm_var_types (Var (x, T)) = |
54789 | 462 |
Var (x, (case AList.lookup (op =) normTs x of NONE => T | SOME T' => T')) |
52758
7ffcd6f2890d
avoid duplicating Var when the types do not quite fit -- since this step occurs before type inference
blanchet
parents:
52125
diff
changeset
|
463 |
| norm_var_types t = t |
7ffcd6f2890d
avoid duplicating Var when the types do not quite fit -- since this step occurs before type inference
blanchet
parents:
52125
diff
changeset
|
464 |
in t |> map_aterms norm_var_types |> fold_rev quant_of (map Var normTs) end |
37991 | 465 |
|
54811 | 466 |
(* Interpret an ATP formula as a HOL term, extracting sort constraints as they appear in the |
467 |
formula. *) |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
468 |
fun prop_of_atp ctxt format type_enc textual sym_tab phi = |
38014 | 469 |
let |
470 |
fun do_formula pos phi = |
|
54789 | 471 |
(case phi of |
38014 | 472 |
AQuant (_, [], phi) => do_formula pos phi |
42526 | 473 |
| AQuant (q, (s, _) :: xs, phi') => |
38014 | 474 |
do_formula pos (AQuant (q, xs, phi')) |
42526 | 475 |
(* FIXME: TFF *) |
54768
ee0881a54c72
fixed confusion between 'prop' and 'bool' introduced in 4960647932ec
blanchet
parents:
54757
diff
changeset
|
476 |
#>> quantify_over_var (case q of AForall => forall_of | AExists => exists_of) |
55185
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
477 |
(repair_var_name textual s) |
38014 | 478 |
| AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not |
37991 | 479 |
| AConn (c, [phi1, phi2]) => |
38014 | 480 |
do_formula (pos |> c = AImplies ? not) phi1 |
481 |
##>> do_formula pos phi2 |
|
482 |
#>> (case c of |
|
54811 | 483 |
AAnd => s_conj |
484 |
| AOr => s_disj |
|
485 |
| AImplies => s_imp |
|
486 |
| AIff => s_iff |
|
487 |
| ANot => raise Fail "impossible connective") |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
488 |
| AAtom tm => term_of_atom ctxt format type_enc textual sym_tab pos tm |
54789 | 489 |
| _ => raise ATP_FORMULA [phi]) |
490 |
in |
|
491 |
repair_tvar_sorts (do_formula true phi Vartab.empty) |
|
492 |
end |
|
37991 | 493 |
|
54500 | 494 |
fun find_first_in_list_vector vec key = |
495 |
Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key |
|
496 |
| (_, value) => value) NONE vec |
|
497 |
||
498 |
val unprefix_fact_number = space_implode "_" o tl o space_explode "_" |
|
499 |
||
500 |
fun resolve_one_named_fact fact_names s = |
|
54789 | 501 |
(case try (unprefix fact_prefix) s of |
54500 | 502 |
SOME s' => |
503 |
let val s' = s' |> unprefix_fact_number |> unascii_of in |
|
504 |
s' |> find_first_in_list_vector fact_names |> Option.map (pair s') |
|
505 |
end |
|
54789 | 506 |
| NONE => NONE) |
54506 | 507 |
|
54500 | 508 |
fun resolve_fact fact_names = map_filter (resolve_one_named_fact fact_names) |
509 |
||
510 |
fun resolve_one_named_conjecture s = |
|
54789 | 511 |
(case try (unprefix conjecture_prefix) s of |
54500 | 512 |
SOME s' => Int.fromString s' |
54789 | 513 |
| NONE => NONE) |
54500 | 514 |
|
515 |
val resolve_conjecture = map_filter resolve_one_named_conjecture |
|
516 |
||
517 |
fun is_axiom_used_in_proof pred = |
|
518 |
exists (fn ((_, ss), _, _, _, []) => exists pred ss | _ => false) |
|
519 |
||
520 |
fun add_non_rec_defs fact_names accum = |
|
521 |
Vector.foldl (fn (facts, facts') => |
|
54818 | 522 |
union (op =) (filter (fn (_, (_, status)) => status = Non_Rec_Def) facts) facts') |
54500 | 523 |
accum fact_names |
524 |
||
525 |
val leo2_extcnf_equal_neg_rule = "extcnf_equal_neg" |
|
526 |
val leo2_unfold_def_rule = "unfold_def" |
|
527 |
||
528 |
fun add_fact ctxt fact_names ((_, ss), _, _, rule, deps) = |
|
529 |
(if rule = leo2_extcnf_equal_neg_rule then |
|
56104
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents:
55285
diff
changeset
|
530 |
insert (op =) (short_thm_name ctxt ext, (Global, General)) |
54500 | 531 |
else if rule = leo2_unfold_def_rule then |
54789 | 532 |
(* LEO 1.3.3 does not record definitions properly, leading to missing dependencies in the TSTP |
533 |
proof. Remove the next line once this is fixed. *) |
|
54500 | 534 |
add_non_rec_defs fact_names |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54774
diff
changeset
|
535 |
else if rule = agsyhol_core_rule orelse rule = satallax_core_rule then |
54500 | 536 |
(fn [] => |
54789 | 537 |
(* agsyHOL and Satallax don't include definitions in their unsatisfiable cores, so we |
538 |
assume the worst and include them all here. *) |
|
56104
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents:
55285
diff
changeset
|
539 |
[(short_thm_name ctxt ext, (Global, General))] |> add_non_rec_defs fact_names |
54500 | 540 |
| facts => facts) |
541 |
else |
|
542 |
I) |
|
543 |
#> (if null deps then union (op =) (resolve_fact fact_names ss) else I) |
|
544 |
||
545 |
fun used_facts_in_atp_proof ctxt fact_names atp_proof = |
|
546 |
if null atp_proof then Vector.foldl (uncurry (union (op =))) [] fact_names |
|
547 |
else fold (add_fact ctxt fact_names) atp_proof [] |
|
548 |
||
549 |
fun used_facts_in_unsound_atp_proof _ _ [] = NONE |
|
550 |
| used_facts_in_unsound_atp_proof ctxt fact_names atp_proof = |
|
551 |
let val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof in |
|
552 |
if forall (fn (_, (sc, _)) => sc = Global) used_facts andalso |
|
54506 | 553 |
not (is_axiom_used_in_proof (not o null o resolve_conjecture o single) atp_proof) then |
54500 | 554 |
SOME (map fst used_facts) |
555 |
else |
|
556 |
NONE |
|
557 |
end |
|
558 |
||
559 |
val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix |
|
560 |
||
561 |
(* overapproximation (good enough) *) |
|
562 |
fun is_lam_lifted s = |
|
563 |
String.isPrefix fact_prefix s andalso |
|
564 |
String.isSubstring ascii_of_lam_fact_prefix s |
|
565 |
||
566 |
val is_combinator_def = String.isPrefix (helper_prefix ^ combinator_prefix) |
|
567 |
||
55257 | 568 |
fun atp_proof_prefers_lifting atp_proof = |
569 |
(is_axiom_used_in_proof is_combinator_def atp_proof, |
|
570 |
is_axiom_used_in_proof is_lam_lifted atp_proof) = (false, true) |
|
54500 | 571 |
|
572 |
val is_typed_helper_name = |
|
573 |
String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix |
|
574 |
||
575 |
fun is_typed_helper_used_in_atp_proof atp_proof = |
|
576 |
is_axiom_used_in_proof is_typed_helper_name atp_proof |
|
577 |
||
54772 | 578 |
fun replace_one_dependency (old, new) dep = if is_same_atp_step dep old then new else [dep] |
579 |
fun replace_dependencies_in_line old_new (name, role, t, rule, deps) = |
|
580 |
(name, role, t, rule, fold (union (op =) o replace_one_dependency old_new) deps []) |
|
581 |
||
54499 | 582 |
fun repair_name "$true" = "c_True" |
583 |
| repair_name "$false" = "c_False" |
|
584 |
| repair_name "$$e" = tptp_equal (* seen in Vampire proofs *) |
|
585 |
| repair_name s = |
|
586 |
if is_tptp_equal s orelse |
|
587 |
(* seen in Vampire proofs *) |
|
588 |
(String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then |
|
589 |
tptp_equal |
|
590 |
else |
|
591 |
s |
|
592 |
||
593 |
fun infer_formula_types ctxt = |
|
594 |
Type.constraint HOLogic.boolT |
|
54756 | 595 |
#> Syntax.check_term (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) |
54499 | 596 |
|
597 |
val combinator_table = |
|
598 |
[(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def [abs_def]}), |
|
599 |
(@{const_name Meson.COMBK}, @{thm Meson.COMBK_def [abs_def]}), |
|
600 |
(@{const_name Meson.COMBB}, @{thm Meson.COMBB_def [abs_def]}), |
|
601 |
(@{const_name Meson.COMBC}, @{thm Meson.COMBC_def [abs_def]}), |
|
602 |
(@{const_name Meson.COMBS}, @{thm Meson.COMBS_def [abs_def]})] |
|
603 |
||
604 |
fun uncombine_term thy = |
|
605 |
let |
|
606 |
fun aux (t1 $ t2) = betapply (pairself aux (t1, t2)) |
|
607 |
| aux (Abs (s, T, t')) = Abs (s, T, aux t') |
|
608 |
| aux (t as Const (x as (s, _))) = |
|
609 |
(case AList.lookup (op =) combinator_table s of |
|
54756 | 610 |
SOME thm => thm |> prop_of |> specialize_type thy x |> Logic.dest_equals |> snd |
611 |
| NONE => t) |
|
54499 | 612 |
| aux t = t |
613 |
in aux end |
|
614 |
||
615 |
fun unlift_term lifted = |
|
616 |
map_aterms (fn t as Const (s, _) => |
|
617 |
if String.isPrefix lam_lifted_prefix s then |
|
54756 | 618 |
(* FIXME: do something about the types *) |
619 |
(case AList.lookup (op =) lifted s of |
|
620 |
SOME t => unlift_term lifted t |
|
621 |
| NONE => t) |
|
54499 | 622 |
else |
623 |
t |
|
624 |
| t => t) |
|
625 |
||
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
626 |
fun termify_line _ _ _ _ _ (_, Type_Role, _, _, _) = NONE |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
627 |
| termify_line ctxt format type_enc lifted sym_tab (name, role, u, rule, deps) = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
628 |
let |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
629 |
val thy = Proof_Context.theory_of ctxt |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
630 |
val t = u |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
631 |
|> prop_of_atp ctxt format type_enc true sym_tab |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
632 |
|> uncombine_term thy |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
633 |
|> unlift_term lifted |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
634 |
|> infer_formula_types ctxt |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
635 |
|> HOLogic.mk_Trueprop |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
636 |
in |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
637 |
SOME (name, role, t, rule, deps) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
638 |
end |
54499 | 639 |
|
640 |
val waldmeister_conjecture_num = "1.0.0.0" |
|
641 |
||
54756 | 642 |
fun repair_waldmeister_endgame proof = |
54499 | 643 |
let |
56854 | 644 |
fun repair_tail (name, _, @{const Trueprop} $ t, rule, deps) = |
645 |
(name, Negated_Conjecture, @{const Trueprop} $ s_not t, rule, deps) |
|
54756 | 646 |
fun repair_body [] = [] |
647 |
| repair_body ((line as ((num, _), _, _, _, _)) :: lines) = |
|
648 |
if num = waldmeister_conjecture_num then map repair_tail (line :: lines) |
|
649 |
else line :: repair_body lines |
|
650 |
in |
|
651 |
repair_body proof |
|
652 |
end |
|
54499 | 653 |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
654 |
fun termify_atp_proof ctxt prover format type_enc pool lifted sym_tab = |
54499 | 655 |
clean_up_atp_proof_dependencies |
656 |
#> nasty_atp_proof pool |
|
657 |
#> map_term_names_in_atp_proof repair_name |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
658 |
#> map_filter (termify_line ctxt format type_enc lifted sym_tab) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
659 |
#> prover = waldmeisterN ? repair_waldmeister_endgame |
54499 | 660 |
|
55192
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
661 |
fun take_distinct_vars seen ((t as Var _) :: ts) = |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
662 |
if member (op aconv) seen t then rev seen else take_distinct_vars (t :: seen) ts |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
663 |
| take_distinct_vars seen _ = rev seen |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
664 |
|
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
665 |
fun unskolemize_term skos t = |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
666 |
let |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
667 |
val is_sko = member (op =) skos |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
668 |
|
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
669 |
fun find_args args (t $ u) = find_args (u :: args) t #> find_args [] u |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
670 |
| find_args _ (Abs (_, _, t)) = find_args [] t |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
671 |
| find_args args (Free (s, _)) = |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
672 |
if is_sko s then |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
673 |
let val new = take_distinct_vars [] args in |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
674 |
(fn [] => new | old => if length new < length old then new else old) |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
675 |
end |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
676 |
else |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
677 |
I |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
678 |
| find_args _ _ = I |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
679 |
|
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
680 |
val alls = find_args [] t [] |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
681 |
val num_alls = length alls |
55195 | 682 |
|
683 |
fun fix_skos args (t $ u) = fix_skos (fix_skos [] u :: args) t |
|
684 |
| fix_skos args (t as Free (s, T)) = |
|
685 |
if is_sko s then list_comb (Free (s, funpow num_alls body_type T), drop num_alls args) |
|
686 |
else list_comb (t, args) |
|
687 |
| fix_skos [] (Abs (s, T, t)) = Abs (s, T, fix_skos [] t) |
|
688 |
| fix_skos [] t = t |
|
689 |
| fix_skos args t = list_comb (fix_skos [] t, args) |
|
690 |
||
691 |
val t' = fix_skos [] t |
|
55192
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
692 |
|
55195 | 693 |
fun add_sko (t as Free (s, _)) = is_sko s ? insert (op aconv) t |
694 |
| add_sko _ = I |
|
55192
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
695 |
|
55195 | 696 |
val exs = Term.fold_aterms add_sko t' [] |
697 |
in |
|
698 |
t' |
|
699 |
|> HOLogic.dest_Trueprop |
|
700 |
|> fold exists_of exs |> Term.map_abs_vars (K Name.uu) |
|
701 |
|> fold_rev forall_of alls |
|
702 |
|> HOLogic.mk_Trueprop |
|
55192
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
703 |
end |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
704 |
|
54772 | 705 |
fun introduce_spass_skolem [] = [] |
706 |
| introduce_spass_skolem (proof as (_, _, _, rule1, _) :: _) = |
|
707 |
if rule1 = spass_input_rule then |
|
708 |
let |
|
709 |
fun add_sko (Free (s, _)) = String.isPrefix spass_skolem_prefix s ? insert (op =) s |
|
710 |
| add_sko _ = I |
|
711 |
||
712 |
(* union-find would be faster *) |
|
54799 | 713 |
fun add_cycle [] = I |
714 |
| add_cycle ss = |
|
54772 | 715 |
fold (fn s => Graph.default_node (s, ())) ss |
716 |
#> fold Graph.add_edge (ss ~~ tl ss @ [hd ss]) |
|
717 |
||
718 |
val (input_steps, other_steps) = List.partition (null o #5) proof |
|
719 |
||
720 |
val skoss = map (fn (_, _, t, _, _) => Term.fold_aterms add_sko t []) input_steps |
|
721 |
val skoss_input_steps = filter_out (null o fst) (skoss ~~ input_steps) |
|
54799 | 722 |
val groups = Graph.strong_conn (fold add_cycle skoss Graph.empty) |
54772 | 723 |
|
724 |
fun step_name_of_group skos = (implode skos, []) |
|
725 |
fun in_group group = member (op =) group o hd |
|
726 |
fun group_of sko = the (find_first (fn group => in_group group sko) groups) |
|
727 |
||
55192
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
728 |
fun new_steps (skoss_steps : (string list * (term, 'a) atp_step) list) group = |
54772 | 729 |
let |
55192
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
730 |
val name = step_name_of_group group |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
731 |
val name0 = name |>> prefix "0" |
54772 | 732 |
val t = |
733 |
skoss_steps |
|
734 |
|> map (snd #> #3 #> HOLogic.dest_Trueprop) |
|
735 |
|> Library.foldr1 s_conj |
|
736 |
|> HOLogic.mk_Trueprop |
|
55192
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
737 |
val skos = fold (union (op =) o fst) skoss_steps [] |
54772 | 738 |
val deps = map (snd #> #1) skoss_steps |
739 |
in |
|
55195 | 740 |
[(name0, Unknown, unskolemize_term skos t, spass_pre_skolemize_rule, deps), |
741 |
(name, Unknown, t, spass_skolemize_rule, [name0])] |
|
54772 | 742 |
end |
743 |
||
744 |
val sko_steps = |
|
55192
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
745 |
maps (fn group => new_steps (filter (in_group group o fst) skoss_input_steps) group) |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
746 |
groups |
54772 | 747 |
|
748 |
val old_news = |
|
749 |
map (fn (skos, (name, _, _, _, _)) => (name, [step_name_of_group (group_of skos)])) |
|
750 |
skoss_input_steps |
|
751 |
val repair_deps = fold replace_dependencies_in_line old_news |
|
752 |
in |
|
753 |
input_steps @ sko_steps @ map repair_deps other_steps |
|
754 |
end |
|
755 |
else |
|
756 |
proof |
|
757 |
||
54505 | 758 |
fun factify_atp_proof fact_names hyp_ts concl_t atp_proof = |
759 |
let |
|
54799 | 760 |
fun factify_step ((num, ss), _, t, rule, deps) = |
54505 | 761 |
let |
762 |
val (ss', role', t') = |
|
763 |
(case resolve_conjecture ss of |
|
764 |
[j] => |
|
765 |
if j = length hyp_ts then ([], Conjecture, concl_t) else ([], Hypothesis, nth hyp_ts j) |
|
55185
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
766 |
| _ => |
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
767 |
(case resolve_fact fact_names ss of |
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
768 |
[] => (ss, Plain, t) |
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
769 |
| facts => (map fst facts, Axiom, t))) |
54505 | 770 |
in |
771 |
((num, ss'), role', t', rule, deps) |
|
772 |
end |
|
773 |
||
774 |
val atp_proof = map factify_step atp_proof |
|
775 |
val names = map #1 atp_proof |
|
776 |
||
777 |
fun repair_dep (num, ss) = (num, the_default ss (AList.lookup (op =) names num)) |
|
778 |
fun repair_deps (name, role, t, rule, deps) = (name, role, t, rule, map repair_dep deps) |
|
54772 | 779 |
in |
780 |
map repair_deps atp_proof |
|
781 |
end |
|
54505 | 782 |
|
31038 | 783 |
end; |