author | blanchet |
Mon, 25 Sep 2023 17:16:32 +0200 | |
changeset 78696 | ef89f1beee95 |
parent 78693 | 1c0576840bf4 |
child 78697 | 8ca71c0ae31f |
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 |
|
58653
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
blanchet
parents:
58652
diff
changeset
|
31 |
val leo2_extcnf_equal_neg_rule : string |
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
blanchet
parents:
58652
diff
changeset
|
32 |
val satallax_tab_rule_prefix : string |
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
blanchet
parents:
58652
diff
changeset
|
33 |
|
49983
33e18e9916a8
use metaquantification when possible in Isar proofs
blanchet
parents:
49914
diff
changeset
|
34 |
val forall_of : term -> term -> term |
33e18e9916a8
use metaquantification when possible in Isar proofs
blanchet
parents:
49914
diff
changeset
|
35 |
val exists_of : term -> term -> term |
57767
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
blanchet
parents:
57765
diff
changeset
|
36 |
val simplify_bool : term -> term |
57792 | 37 |
val var_name_of_typ : typ -> string |
57765 | 38 |
val rename_bound_vars : term -> term |
55285 | 39 |
val type_enc_aliases : (string * string list) list |
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
40 |
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
|
41 |
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
|
42 |
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
|
43 |
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
|
44 |
bool -> int Symtab.table -> |
54811 | 45 |
(string, string, (string, string atp_type) atp_term, string) atp_formula -> term |
54499 | 46 |
|
77432 | 47 |
val is_conjecture_used_in_proof : string atp_proof -> bool |
57263 | 48 |
val used_facts_in_atp_proof : Proof.context -> (string * stature) list -> string atp_proof -> |
49 |
(string * stature) list |
|
55257 | 50 |
val atp_proof_prefers_lifting : string atp_proof -> bool |
54500 | 51 |
val is_typed_helper_used_in_atp_proof : string atp_proof -> bool |
54772 | 52 |
val replace_dependencies_in_line : atp_step_name * atp_step_name list -> ('a, 'b) atp_step -> |
53 |
('a, 'b) atp_step |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
54 |
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
|
55 |
ATP_Problem_Generate.type_enc -> string Symtab.table -> (string * term) list -> |
54499 | 56 |
int Symtab.table -> string atp_proof -> (term, string) atp_step list |
57267
8b87114357bd
integrated new Waldmeister code with 'sledgehammer' command
blanchet
parents:
57263
diff
changeset
|
57 |
val repair_waldmeister_endgame : (term, 'a) atp_step list -> (term, 'a) atp_step list |
57699 | 58 |
val infer_formulas_types : Proof.context -> term list -> term list |
72401
2783779b7dd3
removed obsolete unmaintained experimental prover Pirate
blanchet
parents:
70940
diff
changeset
|
59 |
val introduce_spass_skolems : (term, string) atp_step list -> (term, string) atp_step list |
57263 | 60 |
val factify_atp_proof : (string * 'a) list -> term list -> term -> (term, string) atp_step list -> |
61 |
(term, string) atp_step list |
|
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
62 |
val termify_atp_abduce_candidate : Proof.context -> string -> ATP_Problem.atp_format -> |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
63 |
ATP_Problem_Generate.type_enc -> string Symtab.table -> (string * term) list -> |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
64 |
int Symtab.table -> (string, string, (string, string atp_type) atp_term, string) atp_formula -> |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
65 |
term |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
66 |
val top_abduce_candidates : int -> term list -> term list |
77602
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
67 |
val sort_propositions_by_provability : Proof.context -> term list -> term list |
24425
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
68 |
end; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
69 |
|
46320 | 70 |
structure ATP_Proof_Reconstruct : ATP_PROOF_RECONSTRUCT = |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
71 |
struct |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
72 |
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43062
diff
changeset
|
73 |
open ATP_Util |
38028 | 74 |
open ATP_Problem |
39452 | 75 |
open ATP_Proof |
46320 | 76 |
open ATP_Problem_Generate |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
77 |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
78 |
val metisN = "metis" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
79 |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
80 |
val full_typesN = "full_types" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
81 |
val partial_typesN = "partial_types" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
82 |
val no_typesN = "no_types" |
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 |
val really_full_type_enc = "mono_tags" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
85 |
val full_type_enc = "poly_guards_query" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
86 |
val partial_type_enc = "poly_args" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
87 |
val no_type_enc = "erased" |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
88 |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
89 |
val full_type_encs = [full_type_enc, really_full_type_enc] |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
90 |
val partial_type_encs = partial_type_enc :: full_type_encs |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
91 |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
92 |
val type_enc_aliases = |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
93 |
[(full_typesN, full_type_encs), |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
94 |
(partial_typesN, partial_type_encs), |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
95 |
(no_typesN, [no_type_enc])] |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
96 |
|
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
97 |
fun unalias_type_enc s = |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
98 |
AList.lookup (op =) type_enc_aliases s |> the_default [s] |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
99 |
|
54500 | 100 |
val default_metis_lam_trans = combsN |
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45511
diff
changeset
|
101 |
|
58653
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
blanchet
parents:
58652
diff
changeset
|
102 |
val leo2_extcnf_equal_neg_rule = "extcnf_equal_neg" |
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
blanchet
parents:
58652
diff
changeset
|
103 |
val satallax_tab_rule_prefix = "tab_" |
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
blanchet
parents:
58652
diff
changeset
|
104 |
|
50670
eaa540986291
properly take the existential closure of skolems
blanchet
parents:
49983
diff
changeset
|
105 |
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
|
106 |
| term_name' _ = "" |
50670
eaa540986291
properly take the existential closure of skolems
blanchet
parents:
49983
diff
changeset
|
107 |
|
eaa540986291
properly take the existential closure of skolems
blanchet
parents:
49983
diff
changeset
|
108 |
fun lambda' v = Term.lambda_name (term_name' v, v) |
eaa540986291
properly take the existential closure of skolems
blanchet
parents:
49983
diff
changeset
|
109 |
|
78693
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
110 |
fun If_const T = Const (\<^const_name>\<open>If\<close>, HOLogic.boolT --> T --> T --> T) |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
111 |
|
50670
eaa540986291
properly take the existential closure of skolems
blanchet
parents:
49983
diff
changeset
|
112 |
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
|
113 |
fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda' v t |
39425 | 114 |
|
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
|
115 |
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
|
116 |
let val ww = "'" ^ w in |
69593 | 117 |
TFree (ww, the_default \<^sort>\<open>type\<close> (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
|
118 |
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
|
119 |
|
69593 | 120 |
fun simplify_bool ((all as Const (\<^const_name>\<open>All\<close>, _)) $ Abs (s, T, t)) = |
57767
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
blanchet
parents:
57765
diff
changeset
|
121 |
let val t' = simplify_bool t in |
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
blanchet
parents:
57765
diff
changeset
|
122 |
if loose_bvar1 (t', 0) then all $ Abs (s, T, t') else t' |
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
blanchet
parents:
57765
diff
changeset
|
123 |
end |
69593 | 124 |
| simplify_bool (Const (\<^const_name>\<open>Not\<close>, _) $ t) = s_not (simplify_bool t) |
125 |
| simplify_bool (Const (\<^const_name>\<open>conj\<close>, _) $ t $ u) = |
|
57768
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents:
57767
diff
changeset
|
126 |
s_conj (simplify_bool t, simplify_bool u) |
69593 | 127 |
| simplify_bool (Const (\<^const_name>\<open>disj\<close>, _) $ t $ u) = |
57768
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents:
57767
diff
changeset
|
128 |
s_disj (simplify_bool t, simplify_bool u) |
69593 | 129 |
| simplify_bool (Const (\<^const_name>\<open>implies\<close>, _) $ t $ u) = |
57768
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents:
57767
diff
changeset
|
130 |
s_imp (simplify_bool t, simplify_bool u) |
69593 | 131 |
| simplify_bool ((t as Const (\<^const_name>\<open>HOL.eq\<close>, _)) $ u $ v) = |
57768
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents:
57767
diff
changeset
|
132 |
(case (u, v) of |
69593 | 133 |
(Const (\<^const_name>\<open>True\<close>, _), _) => v |
134 |
| (u, Const (\<^const_name>\<open>True\<close>, _)) => u |
|
135 |
| (Const (\<^const_name>\<open>False\<close>, _), v) => s_not v |
|
136 |
| (u, Const (\<^const_name>\<open>False\<close>, _)) => s_not u |
|
74379 | 137 |
| _ => if u aconv v then \<^Const>\<open>True\<close> else t $ simplify_bool u $ simplify_bool v) |
57767
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
blanchet
parents:
57765
diff
changeset
|
138 |
| simplify_bool (t $ u) = simplify_bool t $ simplify_bool u |
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
blanchet
parents:
57765
diff
changeset
|
139 |
| simplify_bool (Abs (s, T, t)) = Abs (s, T, simplify_bool t) |
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
blanchet
parents:
57765
diff
changeset
|
140 |
| simplify_bool t = t |
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
blanchet
parents:
57765
diff
changeset
|
141 |
|
57765 | 142 |
fun single_letter upper s = |
143 |
let val s' = if String.isPrefix "o" s orelse String.isPrefix "O" s then "z" else s in |
|
144 |
String.extract (Name.desymbolize (SOME upper) (Long_Name.base_name s'), 0, SOME 1) |
|
145 |
end |
|
146 |
||
69593 | 147 |
fun var_name_of_typ (Type (\<^type_name>\<open>fun\<close>, [_, T])) = |
57765 | 148 |
if body_type T = HOLogic.boolT then "p" else "f" |
69593 | 149 |
| var_name_of_typ (Type (\<^type_name>\<open>set\<close>, [T])) = |
58600 | 150 |
let fun default () = single_letter true (var_name_of_typ T) in |
58599 | 151 |
(case T of |
152 |
Type (s, [T1, T2]) => if String.isSuffix "prod" s andalso T1 = T2 then "r" else default () |
|
153 |
| _ => default ()) |
|
154 |
end |
|
57765 | 155 |
| var_name_of_typ (Type (s, Ts)) = |
156 |
if String.isSuffix "list" s then var_name_of_typ (the_single Ts) ^ "s" |
|
157 |
else single_letter false (Long_Name.base_name s) |
|
158 |
| var_name_of_typ (TFree (s, _)) = single_letter false (perhaps (try (unprefix "'")) s) |
|
159 |
| var_name_of_typ (TVar ((s, _), T)) = var_name_of_typ (TFree (s, T)) |
|
160 |
||
161 |
fun rename_bound_vars (t $ u) = rename_bound_vars t $ rename_bound_vars u |
|
162 |
| rename_bound_vars (Abs (_, T, t)) = Abs (var_name_of_typ T, T, rename_bound_vars t) |
|
163 |
| rename_bound_vars t = t |
|
164 |
||
54820 | 165 |
exception ATP_CLASS of string list |
54818 | 166 |
exception ATP_TYPE of string atp_type list |
54811 | 167 |
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
|
168 |
exception ATP_FORMULA of |
54811 | 169 |
(string, string, (string, string atp_type) atp_term, string) atp_formula list |
37991 | 170 |
exception SAME of unit |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
171 |
|
54820 | 172 |
fun class_of_atp_class cls = |
173 |
(case unprefix_and_unascii class_prefix cls of |
|
174 |
SOME s => s |
|
175 |
| NONE => raise ATP_CLASS [cls]) |
|
176 |
||
57697
44341963ade3
correctly translate THF functions from terms to types
blanchet
parents:
57635
diff
changeset
|
177 |
fun atp_type_of_atp_term (ATerm ((s, _), us)) = |
44341963ade3
correctly translate THF functions from terms to types
blanchet
parents:
57635
diff
changeset
|
178 |
let val tys = map atp_type_of_atp_term us in |
44341963ade3
correctly translate THF functions from terms to types
blanchet
parents:
57635
diff
changeset
|
179 |
if s = tptp_fun_type then |
44341963ade3
correctly translate THF functions from terms to types
blanchet
parents:
57635
diff
changeset
|
180 |
(case tys of |
44341963ade3
correctly translate THF functions from terms to types
blanchet
parents:
57635
diff
changeset
|
181 |
[ty1, ty2] => AFun (ty1, ty2) |
44341963ade3
correctly translate THF functions from terms to types
blanchet
parents:
57635
diff
changeset
|
182 |
| _ => raise ATP_TYPE tys) |
44341963ade3
correctly translate THF functions from terms to types
blanchet
parents:
57635
diff
changeset
|
183 |
else |
44341963ade3
correctly translate THF functions from terms to types
blanchet
parents:
57635
diff
changeset
|
184 |
AType ((s, []), tys) |
37991 | 185 |
end |
54818 | 186 |
|
58477 | 187 |
(* Type variables are given the basic sort "HOL.type". Some will later be constrained by information |
188 |
from type literals, or by type inference. *) |
|
189 |
fun typ_of_atp_type ctxt (ty as AType ((a, clss), tys)) = |
|
190 |
let val Ts = map (typ_of_atp_type ctxt) tys in |
|
191 |
(case unprefix_and_unascii native_type_prefix a of |
|
192 |
SOME b => typ_of_atp_type ctxt (atp_type_of_atp_term (unmangled_type b)) |
|
193 |
| NONE => |
|
194 |
(case unprefix_and_unascii type_const_prefix a of |
|
195 |
SOME b => Type (invert_const b, Ts) |
|
196 |
| NONE => |
|
197 |
if not (null tys) then |
|
198 |
raise ATP_TYPE [ty] (* only "tconst"s have type arguments *) |
|
199 |
else |
|
200 |
(case unprefix_and_unascii tfree_prefix a of |
|
201 |
SOME b => make_tfree ctxt b |
|
202 |
| NONE => |
|
203 |
(* The term could be an Isabelle variable or a variable from the ATP, say "X1" or "_5018". |
|
204 |
Sometimes variables from the ATP are indistinguishable from Isabelle variables, which |
|
205 |
forces us to use a type parameter in all cases. *) |
|
206 |
Type_Infer.param 0 ("'" ^ perhaps (unprefix_and_unascii tvar_prefix) a, |
|
69593 | 207 |
(if null clss then \<^sort>\<open>type\<close> else map class_of_atp_class clss))))) |
58477 | 208 |
end |
209 |
| typ_of_atp_type ctxt (AFun (ty1, ty2)) = typ_of_atp_type ctxt ty1 --> typ_of_atp_type ctxt ty2 |
|
210 |
||
54818 | 211 |
fun typ_of_atp_term ctxt = typ_of_atp_type ctxt o atp_type_of_atp_term |
212 |
||
54789 | 213 |
(* 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
|
214 |
fun type_constraint_of_term ctxt (u as ATerm ((a, _), us)) = |
54818 | 215 |
(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
|
216 |
(SOME b, [T]) => (b, T) |
54789 | 217 |
| _ => raise ATP_TERM [u]) |
38014 | 218 |
|
43907
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
219 |
(* 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
|
220 |
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
|
221 |
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
|
222 |
| 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
|
223 |
| add_type_constraint _ _ = I |
38014 | 224 |
|
57789
a73255cffb5b
don't rename variables which will be renamed later anyway
blanchet
parents:
57788
diff
changeset
|
225 |
fun repair_var_name s = |
55185
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
226 |
(case unprefix_and_unascii schematic_var_prefix s of |
57789
a73255cffb5b
don't rename variables which will be renamed later anyway
blanchet
parents:
57788
diff
changeset
|
227 |
SOME s' => s' |
a73255cffb5b
don't rename variables which will be renamed later anyway
blanchet
parents:
57788
diff
changeset
|
228 |
| NONE => s) |
55185
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
229 |
|
54789 | 230 |
(* The number of type arguments of a constant, zero if it's monomorphic. For (instances of) Skolem |
231 |
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
|
232 |
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
|
233 |
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
|
234 |
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
|
235 |
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
|
236 |
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
|
237 |
else |
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43863
diff
changeset
|
238 |
(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
|
239 |
|
69593 | 240 |
fun slack_fastype_of t = fastype_of t handle TERM _ => Type_Infer.anyT \<^sort>\<open>type\<close> |
43182
649bada59658
slacker version of "fastype_of", in case a function has dummy type
blanchet
parents:
43176
diff
changeset
|
241 |
|
54772 | 242 |
val spass_skolem_prefix = "sk" (* "skc" or "skf" *) |
243 |
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
|
244 |
|
78693
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
245 |
val zip_internal_ite_prefix = "zip_internal_ite_" |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
246 |
|
57257 | 247 |
fun var_index_of_textual textual = if textual then 0 else 1 |
248 |
||
249 |
fun quantify_over_var textual quant_of var_s var_T t = |
|
250 |
let |
|
57790 | 251 |
val vars = |
252 |
((var_s, var_index_of_textual textual), var_T) :: |
|
253 |
filter (fn ((s, _), _) => s = var_s) (Term.add_vars t []) |
|
57257 | 254 |
val normTs = vars |> AList.group (op =) |> map (apsnd hd) |
255 |
fun norm_var_types (Var (x, T)) = |
|
256 |
Var (x, the_default T (AList.lookup (op =) normTs x)) |
|
257 |
| norm_var_types t = t |
|
258 |
in t |> map_aterms norm_var_types |> fold_rev quant_of (map Var normTs) end |
|
259 |
||
57790 | 260 |
(* This assumes that distinct names are mapped to distinct names by "Variable.variant_frees". This |
261 |
does not hold in general but should hold for ATP-generated Skolem function names, since these end |
|
262 |
with a digit and "variant_frees" appends letters. *) |
|
263 |
fun fresh_up ctxt s = fst (hd (Variable.variant_frees ctxt [] [(s, ())])) |
|
264 |
||
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
265 |
(* Higher-order translation. Variables are typed (although we don't use that information). Lambdas |
57768
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents:
57767
diff
changeset
|
266 |
are typed. The code is similar to "term_of_atp_fo". *) |
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents:
57767
diff
changeset
|
267 |
fun term_of_atp_ho ctxt sym_tab = |
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
268 |
let |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
269 |
val thy = Proof_Context.theory_of ctxt |
57768
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents:
57767
diff
changeset
|
270 |
val var_index = var_index_of_textual true |
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
271 |
|
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
272 |
fun do_term opt_T u = |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
273 |
(case u of |
78696
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
blanchet
parents:
78693
diff
changeset
|
274 |
AAbs (((var, ty), term), us) => |
57257 | 275 |
let |
276 |
val typ = typ_of_atp_type ctxt ty |
|
57789
a73255cffb5b
don't rename variables which will be renamed later anyway
blanchet
parents:
57788
diff
changeset
|
277 |
val var_name = repair_var_name var |
57257 | 278 |
val tm = do_term NONE term |
78696
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
blanchet
parents:
78693
diff
changeset
|
279 |
val lam = quantify_over_var true lambda' var_name typ tm |
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
blanchet
parents:
78693
diff
changeset
|
280 |
val args = map (do_term NONE) us |
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
blanchet
parents:
78693
diff
changeset
|
281 |
in |
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
blanchet
parents:
78693
diff
changeset
|
282 |
list_comb (lam, args) |
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
blanchet
parents:
78693
diff
changeset
|
283 |
end |
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
284 |
| ATerm ((s, tys), us) => |
74050 | 285 |
if s = "" (* special marker generated on parse error *) then |
286 |
error "Isar proof reconstruction failed because the ATP proof contains unparsable \ |
|
287 |
\material" |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
288 |
else if s = tptp_equal then |
69593 | 289 |
list_comb (Const (\<^const_name>\<open>HOL.eq\<close>, Type_Infer.anyT \<^sort>\<open>type\<close>), |
57768
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents:
57767
diff
changeset
|
290 |
map (do_term NONE) us) |
75052 | 291 |
else if s = tptp_not_equal andalso length us = 2 then |
75048
e926618f9474
handle TPTP '!=' more gracefully in Isar proof reconstruction
blanchet
parents:
74379
diff
changeset
|
292 |
\<^const>\<open>HOL.Not\<close> $ list_comb (Const (\<^const_name>\<open>HOL.eq\<close>, Type_Infer.anyT \<^sort>\<open>type\<close>), |
e926618f9474
handle TPTP '!=' more gracefully in Isar proof reconstruction
blanchet
parents:
74379
diff
changeset
|
293 |
map (do_term NONE) us) |
57256 | 294 |
else if not (null us) then |
295 |
let |
|
57820 | 296 |
val args = map (do_term NONE) us |
57256 | 297 |
val opt_T' = SOME (map slack_fastype_of args ---> the_default dummyT opt_T) |
298 |
val func = do_term opt_T' (ATerm ((s, tys), [])) |
|
299 |
in foldl1 (op $) (func :: args) end |
|
300 |
else if s = tptp_or then HOLogic.disj |
|
301 |
else if s = tptp_and then HOLogic.conj |
|
302 |
else if s = tptp_implies then HOLogic.imp |
|
303 |
else if s = tptp_iff orelse s = tptp_equal then HOLogic.eq_const dummyT |
|
75052 | 304 |
else if s = tptp_not_iff orelse s = tptp_not_equal then \<^term>\<open>\<lambda>x y. x \<noteq> y\<close> |
69593 | 305 |
else if s = tptp_if then \<^term>\<open>\<lambda>P Q. Q \<longrightarrow> P\<close> |
306 |
else if s = tptp_not_and then \<^term>\<open>\<lambda>P Q. \<not> (P \<and> Q)\<close> |
|
307 |
else if s = tptp_not_or then \<^term>\<open>\<lambda>P Q. \<not> (P \<or> Q)\<close> |
|
57256 | 308 |
else if s = tptp_not then HOLogic.Not |
309 |
else if s = tptp_ho_forall then HOLogic.all_const dummyT |
|
310 |
else if s = tptp_ho_exists then HOLogic.exists_const dummyT |
|
57709 | 311 |
else if s = tptp_hilbert_choice then HOLogic.choice_const dummyT |
69593 | 312 |
else if s = tptp_hilbert_the then \<^term>\<open>The\<close> |
78693
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
313 |
else if String.isPrefix zip_internal_ite_prefix s then If_const dummyT |
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
314 |
else |
57256 | 315 |
(case unprefix_and_unascii const_prefix s of |
316 |
SOME s' => |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
317 |
let |
57547
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
blanchet
parents:
57267
diff
changeset
|
318 |
val ((s', _), mangled_us) = s' |> unmangled_const |>> `invert_const |
57256 | 319 |
val num_ty_args = length us - the_default 0 (Symtab.lookup sym_tab s) |
320 |
val (type_us, term_us) = chop num_ty_args us |>> append mangled_us |
|
321 |
val term_ts = map (do_term NONE) term_us |
|
322 |
val Ts = map (typ_of_atp_type ctxt) tys @ map (typ_of_atp_term ctxt) type_us |
|
323 |
val T = |
|
324 |
(if not (null Ts) andalso robust_const_num_type_args thy s' = length Ts then |
|
57768
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents:
57767
diff
changeset
|
325 |
try (Sign.const_instance thy) (s', Ts) |
57256 | 326 |
else |
327 |
NONE) |
|
328 |
|> (fn SOME T => T |
|
329 |
| NONE => |
|
330 |
map slack_fastype_of term_ts ---> |
|
69593 | 331 |
the_default (Type_Infer.anyT \<^sort>\<open>type\<close>) opt_T) |
57256 | 332 |
val t = Const (unproxify_const s', T) |
333 |
in list_comb (t, term_ts) end |
|
334 |
| NONE => (* a free or schematic variable *) |
|
335 |
let |
|
336 |
val ts = map (do_term NONE) us |
|
337 |
val T = |
|
58500
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
blanchet
parents:
58484
diff
changeset
|
338 |
(case tys of |
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
blanchet
parents:
58484
diff
changeset
|
339 |
[ty] => typ_of_atp_type ctxt ty |
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
blanchet
parents:
58484
diff
changeset
|
340 |
| _ => |
57256 | 341 |
map slack_fastype_of ts ---> |
58500
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
blanchet
parents:
58484
diff
changeset
|
342 |
(case opt_T of |
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
blanchet
parents:
58484
diff
changeset
|
343 |
SOME T => T |
69593 | 344 |
| NONE => Type_Infer.anyT \<^sort>\<open>type\<close>)) |
57256 | 345 |
val t = |
346 |
(case unprefix_and_unascii fixed_var_prefix s of |
|
347 |
SOME s => Free (s, T) |
|
348 |
| NONE => |
|
57790 | 349 |
if not (is_tptp_variable s) then Free (fresh_up ctxt s, T) |
57789
a73255cffb5b
don't rename variables which will be renamed later anyway
blanchet
parents:
57788
diff
changeset
|
350 |
else Var ((repair_var_name s, var_index), T)) |
57256 | 351 |
in list_comb (t, ts) end)) |
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
352 |
in do_term end |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
353 |
|
56254 | 354 |
(* First-order translation. No types are known for variables. "Type_Infer.anyT @{sort type}" |
355 |
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
|
356 |
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
|
357 |
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
|
358 |
val thy = Proof_Context.theory_of ctxt |
54789 | 359 |
(* For Metis, we use 1 rather than 0 because variable references in clauses may otherwise |
360 |
conflict with variable constraints in the goal. At least, type inference often fails |
|
361 |
otherwise. See also "axiom_inference" in "Metis_Reconstruct". *) |
|
57257 | 362 |
val var_index = var_index_of_textual textual |
57199 | 363 |
|
43131
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents:
43130
diff
changeset
|
364 |
fun do_term extra_ts opt_T u = |
54789 | 365 |
(case u of |
54818 | 366 |
ATerm ((s, tys), us) => |
74050 | 367 |
if s = "" (* special marker generated on parse error *) then |
57199 | 368 |
error "Isar proof reconstruction failed because the ATP proof contains unparsable \ |
63692 | 369 |
\material" |
51878 | 370 |
else if String.isPrefix native_type_prefix s then |
74379 | 371 |
\<^Const>\<open>True\<close> (* ignore TPTP type information (needed?) *) |
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44399
diff
changeset
|
372 |
else if s = tptp_equal then |
69593 | 373 |
list_comb (Const (\<^const_name>\<open>HOL.eq\<close>, Type_Infer.anyT \<^sort>\<open>type\<close>), |
57768
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents:
57767
diff
changeset
|
374 |
map (do_term [] NONE) us) |
54789 | 375 |
else |
376 |
(case unprefix_and_unascii const_prefix s of |
|
377 |
SOME s' => |
|
54818 | 378 |
let val ((s', s''), mangled_us) = s' |> unmangled_const |>> `invert_const in |
54789 | 379 |
if s' = type_tag_name then |
380 |
(case mangled_us @ us of |
|
54818 | 381 |
[typ_u, term_u] => do_term extra_ts (SOME (typ_of_atp_term ctxt typ_u)) term_u |
54789 | 382 |
| _ => raise ATP_TERM us) |
383 |
else if s' = predicator_name then |
|
69593 | 384 |
do_term [] (SOME \<^typ>\<open>bool\<close>) (hd us) |
54789 | 385 |
else if s' = app_op_name then |
386 |
let val extra_t = do_term [] NONE (List.last us) in |
|
387 |
do_term (extra_t :: extra_ts) |
|
57199 | 388 |
(case opt_T of SOME T => SOME (slack_fastype_of extra_t --> T) | NONE => NONE) |
389 |
(nth us (length us - 2)) |
|
54789 | 390 |
end |
391 |
else if s' = type_guard_name then |
|
74379 | 392 |
\<^Const>\<open>True\<close> (* ignore type predicates *) |
54789 | 393 |
else |
394 |
let |
|
395 |
val new_skolem = String.isPrefix new_skolem_const_prefix s'' |
|
54818 | 396 |
val num_ty_args = length us - the_default 0 (Symtab.lookup sym_tab s) |
397 |
val (type_us, term_us) = chop num_ty_args us |>> append mangled_us |
|
54789 | 398 |
val term_ts = map (do_term [] NONE) term_us |
54818 | 399 |
|
400 |
val Ts = map (typ_of_atp_type ctxt) tys @ map (typ_of_atp_term ctxt) type_us |
|
54789 | 401 |
val T = |
54818 | 402 |
(if not (null Ts) andalso robust_const_num_type_args thy s' = length Ts then |
57199 | 403 |
if new_skolem then SOME (Type_Infer.paramify_vars (tl Ts ---> hd Ts)) |
404 |
else if textual then try (Sign.const_instance thy) (s', Ts) |
|
405 |
else NONE |
|
54789 | 406 |
else |
407 |
NONE) |
|
408 |
|> (fn SOME T => T |
|
409 |
| NONE => |
|
56254 | 410 |
map slack_fastype_of term_ts ---> |
69593 | 411 |
the_default (Type_Infer.anyT \<^sort>\<open>type\<close>) opt_T) |
54789 | 412 |
val t = |
413 |
if new_skolem then Var ((new_skolem_var_name_of_const s'', var_index), T) |
|
414 |
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
|
415 |
in |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
416 |
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
|
417 |
end |
54789 | 418 |
end |
419 |
| NONE => (* a free or schematic variable *) |
|
420 |
let |
|
421 |
val term_ts = |
|
422 |
map (do_term [] NONE) us |
|
423 |
(* SPASS (3.8ds) and Vampire (2.6) pass arguments to Skolem functions in reverse |
|
57699 | 424 |
order, which is incompatible with "metis"'s new skolemizer. *) |
54789 | 425 |
|> exists (fn pre => String.isPrefix pre s) |
426 |
[spass_skolem_prefix, vampire_skolem_prefix] ? rev |
|
427 |
val ts = term_ts @ extra_ts |
|
428 |
val T = |
|
58500
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
blanchet
parents:
58484
diff
changeset
|
429 |
(case tys of |
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
blanchet
parents:
58484
diff
changeset
|
430 |
[ty] => typ_of_atp_type ctxt ty |
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
blanchet
parents:
58484
diff
changeset
|
431 |
| _ => |
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
blanchet
parents:
58484
diff
changeset
|
432 |
(case opt_T of |
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
blanchet
parents:
58484
diff
changeset
|
433 |
SOME T => map slack_fastype_of term_ts ---> T |
69593 | 434 |
| NONE => map slack_fastype_of ts ---> Type_Infer.anyT \<^sort>\<open>type\<close>)) |
54789 | 435 |
val t = |
436 |
(case unprefix_and_unascii fixed_var_prefix s of |
|
437 |
SOME s => Free (s, T) |
|
36909
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents:
36607
diff
changeset
|
438 |
| NONE => |
55185
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
439 |
if textual andalso not (is_tptp_variable s) then |
57790 | 440 |
Free (s |> textual ? fresh_up ctxt, T) |
55185
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
441 |
else |
57789
a73255cffb5b
don't rename variables which will be renamed later anyway
blanchet
parents:
57788
diff
changeset
|
442 |
Var ((repair_var_name s, var_index), T)) |
54789 | 443 |
in list_comb (t, ts) end)) |
43093 | 444 |
in do_term [] end |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
445 |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
446 |
fun term_of_atp ctxt (ATP_Problem.THF _) type_enc = |
57768
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents:
57767
diff
changeset
|
447 |
if ATP_Problem_Generate.is_type_enc_higher_order type_enc then K (term_of_atp_ho ctxt) |
63692 | 448 |
else error "Unsupported Isar reconstruction" |
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
449 |
| term_of_atp ctxt _ type_enc = |
57768
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents:
57767
diff
changeset
|
450 |
if not (ATP_Problem_Generate.is_type_enc_higher_order type_enc) then term_of_atp_fo ctxt |
63692 | 451 |
else error "Unsupported Isar reconstruction" |
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
452 |
|
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
453 |
fun term_of_atom ctxt format type_enc textual sym_tab pos (u as ATerm ((s, _), _)) = |
38014 | 454 |
if String.isPrefix class_prefix s then |
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51878
diff
changeset
|
455 |
add_type_constraint pos (type_constraint_of_term ctxt u) |
74379 | 456 |
#> pair \<^Const>\<open>True\<close> |
38014 | 457 |
else |
69593 | 458 |
pair (term_of_atp ctxt format type_enc textual sym_tab (SOME \<^typ>\<open>bool\<close>) u) |
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36396
diff
changeset
|
459 |
|
37991 | 460 |
(* Update schematic type variables with detected sort constraints. It's not |
42563 | 461 |
totally clear whether this code is necessary. *) |
38014 | 462 |
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
|
463 |
let |
37991 | 464 |
fun do_type (Type (a, Ts)) = Type (a, map do_type Ts) |
465 |
| do_type (TVar (xi, s)) = |
|
466 |
TVar (xi, the_default s (Vartab.lookup tvar_tab xi)) |
|
467 |
| do_type (TFree z) = TFree z |
|
468 |
fun do_term (Const (a, T)) = Const (a, do_type T) |
|
469 |
| do_term (Free (a, T)) = Free (a, do_type T) |
|
470 |
| do_term (Var (xi, T)) = Var (xi, do_type T) |
|
471 |
| do_term (t as Bound _) = t |
|
472 |
| do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t) |
|
473 |
| do_term (t1 $ t2) = do_term t1 $ do_term t2 |
|
474 |
in t |> not (Vartab.is_empty tvar_tab) ? do_term end |
|
475 |
||
54811 | 476 |
(* Interpret an ATP formula as a HOL term, extracting sort constraints as they appear in the |
477 |
formula. *) |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
478 |
fun prop_of_atp ctxt format type_enc textual sym_tab phi = |
38014 | 479 |
let |
480 |
fun do_formula pos phi = |
|
54789 | 481 |
(case phi of |
38014 | 482 |
AQuant (_, [], phi) => do_formula pos phi |
42526 | 483 |
| AQuant (q, (s, _) :: xs, phi') => |
38014 | 484 |
do_formula pos (AQuant (q, xs, phi')) |
42526 | 485 |
(* FIXME: TFF *) |
57257 | 486 |
#>> quantify_over_var textual (case q of AForall => forall_of | AExists => exists_of) |
57789
a73255cffb5b
don't rename variables which will be renamed later anyway
blanchet
parents:
57788
diff
changeset
|
487 |
(repair_var_name s) dummyT |
38014 | 488 |
| AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not |
37991 | 489 |
| AConn (c, [phi1, phi2]) => |
38014 | 490 |
do_formula (pos |> c = AImplies ? not) phi1 |
491 |
##>> do_formula pos phi2 |
|
492 |
#>> (case c of |
|
54811 | 493 |
AAnd => s_conj |
494 |
| AOr => s_disj |
|
495 |
| AImplies => s_imp |
|
496 |
| AIff => s_iff |
|
497 |
| ANot => raise Fail "impossible connective") |
|
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
498 |
| AAtom tm => term_of_atom ctxt format type_enc textual sym_tab pos tm |
54789 | 499 |
| _ => raise ATP_FORMULA [phi]) |
500 |
in |
|
501 |
repair_tvar_sorts (do_formula true phi Vartab.empty) |
|
502 |
end |
|
37991 | 503 |
|
54500 | 504 |
val unprefix_fact_number = space_implode "_" o tl o space_explode "_" |
505 |
||
57263 | 506 |
fun resolve_fact facts s = |
54789 | 507 |
(case try (unprefix fact_prefix) s of |
54500 | 508 |
SOME s' => |
509 |
let val s' = s' |> unprefix_fact_number |> unascii_of in |
|
57263 | 510 |
AList.lookup (op =) facts s' |> Option.map (pair s') |
54500 | 511 |
end |
54789 | 512 |
| NONE => NONE) |
54506 | 513 |
|
57263 | 514 |
fun resolve_conjecture s = |
54789 | 515 |
(case try (unprefix conjecture_prefix) s of |
54500 | 516 |
SOME s' => Int.fromString s' |
54789 | 517 |
| NONE => NONE) |
54500 | 518 |
|
57263 | 519 |
fun resolve_facts facts = map_filter (resolve_fact facts) |
520 |
val resolve_conjectures = map_filter resolve_conjecture |
|
54500 | 521 |
|
522 |
fun is_axiom_used_in_proof pred = |
|
523 |
exists (fn ((_, ss), _, _, _, []) => exists pred ss | _ => false) |
|
524 |
||
77432 | 525 |
val is_conjecture_used_in_proof = is_axiom_used_in_proof (is_some o resolve_conjecture) |
526 |
||
70940
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
blanchet
parents:
69597
diff
changeset
|
527 |
fun add_fact ctxt facts ((num, ss), _, _, rule, deps) = |
58653
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
blanchet
parents:
58652
diff
changeset
|
528 |
(if member (op =) [agsyhol_core_rule, leo2_extcnf_equal_neg_rule] rule orelse |
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
blanchet
parents:
58652
diff
changeset
|
529 |
String.isPrefix satallax_tab_rule_prefix 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 |
532 |
I) |
|
70940
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
blanchet
parents:
69597
diff
changeset
|
533 |
#> (if null deps then union (op =) (resolve_facts facts (num :: ss)) else I) |
54500 | 534 |
|
57263 | 535 |
fun used_facts_in_atp_proof ctxt facts atp_proof = |
536 |
if null atp_proof then facts else fold (add_fact ctxt facts) atp_proof [] |
|
54500 | 537 |
|
538 |
val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix |
|
539 |
||
540 |
(* overapproximation (good enough) *) |
|
541 |
fun is_lam_lifted s = |
|
542 |
String.isPrefix fact_prefix s andalso |
|
543 |
String.isSubstring ascii_of_lam_fact_prefix s |
|
544 |
||
545 |
val is_combinator_def = String.isPrefix (helper_prefix ^ combinator_prefix) |
|
546 |
||
55257 | 547 |
fun atp_proof_prefers_lifting atp_proof = |
548 |
(is_axiom_used_in_proof is_combinator_def atp_proof, |
|
549 |
is_axiom_used_in_proof is_lam_lifted atp_proof) = (false, true) |
|
54500 | 550 |
|
551 |
val is_typed_helper_name = |
|
552 |
String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix |
|
553 |
||
554 |
fun is_typed_helper_used_in_atp_proof atp_proof = |
|
555 |
is_axiom_used_in_proof is_typed_helper_name atp_proof |
|
556 |
||
54772 | 557 |
fun replace_one_dependency (old, new) dep = if is_same_atp_step dep old then new else [dep] |
558 |
fun replace_dependencies_in_line old_new (name, role, t, rule, deps) = |
|
559 |
(name, role, t, rule, fold (union (op =) o replace_one_dependency old_new) deps []) |
|
560 |
||
54499 | 561 |
fun repair_name "$true" = "c_True" |
562 |
| repair_name "$false" = "c_False" |
|
563 |
| repair_name "$$e" = tptp_equal (* seen in Vampire proofs *) |
|
564 |
| repair_name s = |
|
565 |
if is_tptp_equal s orelse |
|
566 |
(* seen in Vampire proofs *) |
|
567 |
(String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then |
|
568 |
tptp_equal |
|
569 |
else |
|
570 |
s |
|
571 |
||
57699 | 572 |
fun set_var_index j = map_aterms (fn Var ((s, 0), T) => Var ((s, j), T) | t => t) |
57635
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
blanchet
parents:
57547
diff
changeset
|
573 |
|
57699 | 574 |
fun infer_formulas_types ctxt = |
57635
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
blanchet
parents:
57547
diff
changeset
|
575 |
map_index (uncurry (fn j => set_var_index j #> Type.constraint HOLogic.boolT)) |
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
blanchet
parents:
57547
diff
changeset
|
576 |
#> Syntax.check_terms (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) |
54499 | 577 |
|
578 |
val combinator_table = |
|
69593 | 579 |
[(\<^const_name>\<open>Meson.COMBI\<close>, @{thm Meson.COMBI_def [abs_def]}), |
580 |
(\<^const_name>\<open>Meson.COMBK\<close>, @{thm Meson.COMBK_def [abs_def]}), |
|
581 |
(\<^const_name>\<open>Meson.COMBB\<close>, @{thm Meson.COMBB_def [abs_def]}), |
|
582 |
(\<^const_name>\<open>Meson.COMBC\<close>, @{thm Meson.COMBC_def [abs_def]}), |
|
583 |
(\<^const_name>\<open>Meson.COMBS\<close>, @{thm Meson.COMBS_def [abs_def]})] |
|
54499 | 584 |
|
585 |
fun uncombine_term thy = |
|
586 |
let |
|
57717
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
587 |
fun uncomb (t1 $ t2) = betapply (uncomb t1, uncomb t2) |
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
588 |
| uncomb (Abs (s, T, t)) = Abs (s, T, uncomb t) |
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
589 |
| uncomb (t as Const (x as (s, _))) = |
54499 | 590 |
(case AList.lookup (op =) combinator_table s of |
59582 | 591 |
SOME thm => thm |> Thm.prop_of |> specialize_type thy x |> Logic.dest_equals |> snd |
54756 | 592 |
| NONE => t) |
57717
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
593 |
| uncomb t = t |
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
594 |
in uncomb end |
54499 | 595 |
|
57717
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
596 |
fun unlift_aterm lifted (t as Const (s, _)) = |
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
597 |
if String.isPrefix lam_lifted_prefix s then |
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
598 |
(* FIXME: do something about the types *) |
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
599 |
(case AList.lookup (op =) lifted s of |
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
600 |
SOME t' => unlift_term lifted t' |
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
601 |
| NONE => t) |
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
602 |
else |
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
603 |
t |
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
604 |
| unlift_aterm _ t = t |
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
605 |
and unlift_term lifted = |
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
606 |
map_aterms (unlift_aterm lifted) |
54499 | 607 |
|
57256 | 608 |
fun termify_line _ _ _ _ _ (_, Type_Role, _, _, _) = NONE |
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
609 |
| 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
|
610 |
let |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
611 |
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
|
612 |
val t = u |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
613 |
|> prop_of_atp ctxt format type_enc true sym_tab |
57717
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
blanchet
parents:
57713
diff
changeset
|
614 |
|> unlift_term lifted |
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
615 |
|> uncombine_term thy |
57768
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents:
57767
diff
changeset
|
616 |
|> simplify_bool |
57255
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
617 |
in |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
618 |
SOME (name, role, t, rule, deps) |
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents:
57199
diff
changeset
|
619 |
end |
54499 | 620 |
|
621 |
val waldmeister_conjecture_num = "1.0.0.0" |
|
622 |
||
54756 | 623 |
fun repair_waldmeister_endgame proof = |
54499 | 624 |
let |
74379 | 625 |
fun repair_tail (name, _, \<^Const>\<open>Trueprop for t\<close>, rule, deps) = |
626 |
(name, Negated_Conjecture, \<^Const>\<open>Trueprop for \<open>s_not t\<close>\<close>, rule, deps) |
|
54756 | 627 |
fun repair_body [] = [] |
628 |
| repair_body ((line as ((num, _), _, _, _, _)) :: lines) = |
|
629 |
if num = waldmeister_conjecture_num then map repair_tail (line :: lines) |
|
630 |
else line :: repair_body lines |
|
631 |
in |
|
632 |
repair_body proof |
|
633 |
end |
|
54499 | 634 |
|
57635
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
blanchet
parents:
57547
diff
changeset
|
635 |
fun map_proof_terms f (lines : ('a * 'b * 'c * 'd * 'e) list) = |
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
blanchet
parents:
57547
diff
changeset
|
636 |
map2 (fn c => fn (a, b, _, d, e) => (a, b, c, d, e)) (f (map #3 lines)) lines |
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
blanchet
parents:
57547
diff
changeset
|
637 |
|
57258
67d85a8aa6cc
Moving the remote prefix deleting on Sledgehammer's side
fleury
parents:
57257
diff
changeset
|
638 |
fun termify_atp_proof ctxt local_prover format type_enc pool lifted sym_tab = |
57267
8b87114357bd
integrated new Waldmeister code with 'sledgehammer' command
blanchet
parents:
57263
diff
changeset
|
639 |
nasty_atp_proof pool |
54499 | 640 |
#> 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
|
641 |
#> map_filter (termify_line ctxt format type_enc lifted sym_tab) |
57699 | 642 |
#> map_proof_terms (infer_formulas_types ctxt #> map HOLogic.mk_Trueprop) |
57258
67d85a8aa6cc
Moving the remote prefix deleting on Sledgehammer's side
fleury
parents:
57257
diff
changeset
|
643 |
#> local_prover = waldmeisterN ? repair_waldmeister_endgame |
54499 | 644 |
|
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
645 |
fun unskolemize_spass_term skos = |
55192
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
646 |
let |
58597 | 647 |
val is_skolem_name = member (op =) skos |
55192
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
648 |
|
58091 | 649 |
fun find_argless_skolem (Free _ $ Var _) = NONE |
58597 | 650 |
| find_argless_skolem (Free (x as (s, _))) = if is_skolem_name s then SOME x else NONE |
58091 | 651 |
| find_argless_skolem (t $ u) = |
652 |
(case find_argless_skolem t of NONE => find_argless_skolem u | sk => sk) |
|
653 |
| find_argless_skolem (Abs (_, _, t)) = find_argless_skolem t |
|
654 |
| find_argless_skolem _ = NONE |
|
55192
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
655 |
|
58597 | 656 |
fun find_skolem_arg (Free (s, _) $ Var v) = if is_skolem_name s then SOME v else NONE |
58091 | 657 |
| find_skolem_arg (t $ u) = (case find_skolem_arg t of NONE => find_skolem_arg u | v => v) |
658 |
| find_skolem_arg (Abs (_, _, t)) = find_skolem_arg t |
|
659 |
| find_skolem_arg _ = NONE |
|
660 |
||
58597 | 661 |
fun kill_skolem_arg (t as Free (s, T) $ Var _) = |
662 |
if is_skolem_name s then Free (s, range_type T) else t |
|
663 |
| kill_skolem_arg (t $ u) = kill_skolem_arg t $ kill_skolem_arg u |
|
664 |
| kill_skolem_arg (Abs (s, T, t)) = Abs (s, T, kill_skolem_arg t) |
|
665 |
| kill_skolem_arg t = t |
|
666 |
||
58091 | 667 |
fun find_var (Var v) = SOME v |
668 |
| find_var (t $ u) = (case find_var t of NONE => find_var u | v => v) |
|
669 |
| find_var (Abs (_, _, t)) = find_var t |
|
670 |
| find_var _ = NONE |
|
671 |
||
672 |
val safe_abstract_over = abstract_over o apsnd (incr_boundvars 1) |
|
55192
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
673 |
|
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
674 |
fun unskolem_inner t = |
58091 | 675 |
(case find_argless_skolem t of |
676 |
SOME (x as (s, T)) => |
|
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
677 |
HOLogic.exists_const T $ Abs (s, T, unskolem_inner (safe_abstract_over (Free x, t))) |
58091 | 678 |
| NONE => |
58597 | 679 |
(case find_skolem_arg t of |
58091 | 680 |
SOME (v as ((s, _), T)) => |
681 |
let |
|
682 |
val (haves, have_nots) = |
|
683 |
HOLogic.disjuncts t |
|
58597 | 684 |
|> List.partition (exists_subterm (curry (op =) (Var v))) |
69593 | 685 |
|> apply2 (fn lits => fold (curry s_disj) lits \<^term>\<open>False\<close>) |
58091 | 686 |
in |
687 |
s_disj (HOLogic.all_const T |
|
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
688 |
$ Abs (s, T, unskolem_inner (safe_abstract_over (Var v, kill_skolem_arg haves))), |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
689 |
unskolem_inner have_nots) |
58091 | 690 |
end |
58597 | 691 |
| NONE => |
692 |
(case find_var t of |
|
693 |
SOME (v as ((s, _), T)) => |
|
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
694 |
HOLogic.all_const T $ Abs (s, T, unskolem_inner (safe_abstract_over (Var v, t))) |
58597 | 695 |
| NONE => t))) |
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
696 |
|
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
697 |
fun unskolem_outer (@{const Trueprop} $ t) = @{const Trueprop} $ unskolem_outer t |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
698 |
| unskolem_outer t = unskolem_inner t |
55195 | 699 |
in |
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
700 |
unskolem_outer |
55192
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
701 |
end |
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents:
55185
diff
changeset
|
702 |
|
57788
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
703 |
fun rename_skolem_args t = |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
704 |
let |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
705 |
fun add_skolem_args (Abs (_, _, t)) = add_skolem_args t |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
706 |
| add_skolem_args t = |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
707 |
(case strip_comb t of |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
708 |
(Free (s, _), args as _ :: _) => |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
709 |
if String.isPrefix spass_skolem_prefix s then |
67522 | 710 |
insert (op =) (s, take_prefix is_Var args) |
57788
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
711 |
else |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
712 |
fold add_skolem_args args |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
713 |
| (u, args as _ :: _) => fold add_skolem_args (u :: args) |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
714 |
| _ => I) |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
715 |
|
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
716 |
fun subst_of_skolem (sk, args) = |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
717 |
map_index (fn (j, Var (z, T)) => (z, Var ((sk ^ "_" ^ string_of_int j, 0), T))) args |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
718 |
|
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
719 |
val subst = maps subst_of_skolem (add_skolem_args t []) |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
720 |
in |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
721 |
subst_vars ([], subst) t |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
722 |
end |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
723 |
|
72401
2783779b7dd3
removed obsolete unmaintained experimental prover Pirate
blanchet
parents:
70940
diff
changeset
|
724 |
fun introduce_spass_skolems proof = |
57787 | 725 |
let |
57788
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
726 |
fun add_skolem (Free (s, _)) = String.isPrefix spass_skolem_prefix s ? insert (op =) s |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
727 |
| add_skolem _ = I |
54772 | 728 |
|
57787 | 729 |
(* union-find would be faster *) |
730 |
fun add_cycle [] = I |
|
731 |
| add_cycle ss = |
|
732 |
fold (fn s => Graph.default_node (s, ())) ss |
|
733 |
#> fold Graph.add_edge (ss ~~ tl ss @ [hd ss]) |
|
54772 | 734 |
|
57787 | 735 |
val (input_steps, other_steps) = List.partition (null o #5) proof |
54772 | 736 |
|
58597 | 737 |
(* The names of the formulas are added to the Skolem constants, to ensure that formulas giving |
738 |
rise to several clauses are skolemized together. *) |
|
739 |
val skoXss = map (fn ((_, ss), _, t, _, _) => Term.fold_aterms add_skolem t ss) input_steps |
|
58601
85fa90262807
avoid creating needless skolemization steps for SPASS
blanchet
parents:
58600
diff
changeset
|
740 |
val groups0 = Graph.strong_conn (fold add_cycle skoXss Graph.empty) |
85fa90262807
avoid creating needless skolemization steps for SPASS
blanchet
parents:
58600
diff
changeset
|
741 |
val groups = filter (exists (String.isPrefix spass_skolem_prefix)) groups0 |
85fa90262807
avoid creating needless skolemization steps for SPASS
blanchet
parents:
58600
diff
changeset
|
742 |
|
85fa90262807
avoid creating needless skolemization steps for SPASS
blanchet
parents:
58600
diff
changeset
|
743 |
val skoXss_input_steps = skoXss ~~ input_steps |
54772 | 744 |
|
58597 | 745 |
fun step_name_of_group skoXs = (implode skoXs, []) |
57787 | 746 |
fun in_group group = member (op =) group o hd |
58601
85fa90262807
avoid creating needless skolemization steps for SPASS
blanchet
parents:
58600
diff
changeset
|
747 |
fun group_of skoX = find_first (fn group => in_group group skoX) groups |
54772 | 748 |
|
58652 | 749 |
fun new_steps (skoXss_steps : (string list * (term, 'a) atp_step) list) group = |
57787 | 750 |
let |
751 |
val name = step_name_of_group group |
|
752 |
val name0 = name |>> prefix "0" |
|
57788
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
753 |
val t = |
58597 | 754 |
(case map (snd #> #3) skoXss_steps of |
57788
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
755 |
[t] => t |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
756 |
| ts => ts |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
757 |
|> map (HOLogic.dest_Trueprop #> rename_skolem_args) |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
758 |
|> Library.foldr1 s_conj |
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
blanchet
parents:
57787
diff
changeset
|
759 |
|> HOLogic.mk_Trueprop) |
58597 | 760 |
val skos = |
761 |
fold (union (op =) o filter (String.isPrefix spass_skolem_prefix) o fst) skoXss_steps [] |
|
762 |
val deps = map (snd #> #1) skoXss_steps |
|
57787 | 763 |
in |
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
764 |
[(name0, Unknown, unskolemize_spass_term skos t, spass_pre_skolemize_rule, deps), |
57787 | 765 |
(name, Unknown, t, spass_skolemize_rule, [name0])] |
766 |
end |
|
54772 | 767 |
|
57787 | 768 |
val sko_steps = |
58597 | 769 |
maps (fn group => new_steps (filter (in_group group o fst) skoXss_input_steps) group) groups |
54772 | 770 |
|
57787 | 771 |
val old_news = |
58601
85fa90262807
avoid creating needless skolemization steps for SPASS
blanchet
parents:
58600
diff
changeset
|
772 |
map_filter (fn (skoXs, (name, _, _, _, _)) => |
85fa90262807
avoid creating needless skolemization steps for SPASS
blanchet
parents:
58600
diff
changeset
|
773 |
Option.map (pair name o single o step_name_of_group) (group_of skoXs)) |
58597 | 774 |
skoXss_input_steps |
57787 | 775 |
val repair_deps = fold replace_dependencies_in_line old_news |
776 |
in |
|
777 |
input_steps @ sko_steps @ map repair_deps other_steps |
|
778 |
end |
|
54772 | 779 |
|
78693
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
780 |
val zf_stmt_prefix = "zf_stmt_" |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
781 |
|
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
782 |
fun is_if_True_or_False_axiom true_or_false t = |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
783 |
(case t of |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
784 |
@{const Trueprop} $ |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
785 |
(Const (@{const_name HOL.eq}, _) $ |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
786 |
(Const (@{const_name If}, _) $ cond $ Var _ $ Var _) $ |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
787 |
Var _) => |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
788 |
cond aconv true_or_false |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
789 |
| _ => false) |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
790 |
|
57263 | 791 |
fun factify_atp_proof facts hyp_ts concl_t atp_proof = |
54505 | 792 |
let |
57785 | 793 |
fun factify_step ((num, ss), role, t, rule, deps) = |
54505 | 794 |
let |
795 |
val (ss', role', t') = |
|
57263 | 796 |
(case resolve_conjectures ss of |
54505 | 797 |
[j] => |
58484
b4c0e2b00036
support hypotheses with schematics in Isar proofs
blanchet
parents:
58477
diff
changeset
|
798 |
if j = length hyp_ts then ([], Conjecture, concl_t) |
b4c0e2b00036
support hypotheses with schematics in Isar proofs
blanchet
parents:
58477
diff
changeset
|
799 |
else ([], Hypothesis, close_form (nth hyp_ts j)) |
55185
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
800 |
| _ => |
70940
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
blanchet
parents:
69597
diff
changeset
|
801 |
(case resolve_facts facts (num :: ss) of |
78693
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
802 |
[] => |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
803 |
if role = Axiom andalso String.isPrefix zf_stmt_prefix num |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
804 |
andalso is_if_True_or_False_axiom @{const True} t then |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
805 |
(["if_True"], Axiom, t) |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
806 |
else if role = Axiom andalso String.isPrefix zf_stmt_prefix num |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
807 |
andalso is_if_True_or_False_axiom @{const False} t then |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
808 |
(["if_False"], Axiom, t) |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
809 |
else |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
810 |
(ss, |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
811 |
if role = Definition then Definition |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
812 |
else if role = Lemma then Lemma |
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
blanchet
parents:
77602
diff
changeset
|
813 |
else Plain, t) |
55185
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents:
54822
diff
changeset
|
814 |
| facts => (map fst facts, Axiom, t))) |
54505 | 815 |
in |
816 |
((num, ss'), role', t', rule, deps) |
|
817 |
end |
|
818 |
||
819 |
val atp_proof = map factify_step atp_proof |
|
820 |
val names = map #1 atp_proof |
|
821 |
||
822 |
fun repair_dep (num, ss) = (num, the_default ss (AList.lookup (op =) names num)) |
|
823 |
fun repair_deps (name, role, t, rule, deps) = (name, role, t, rule, map repair_dep deps) |
|
54772 | 824 |
in |
825 |
map repair_deps atp_proof |
|
826 |
end |
|
54505 | 827 |
|
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
828 |
fun termify_atp_abduce_candidate ctxt local_prover format type_enc pool lifted sym_tab phi = |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
829 |
let |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
830 |
val proof = [(("", []), Conjecture, mk_anot phi, "", [])] |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
831 |
val new_proof = termify_atp_proof ctxt local_prover format type_enc pool lifted sym_tab proof |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
832 |
in |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
833 |
(case new_proof of |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
834 |
[(_, _, phi', _, _)] => phi' |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
835 |
| _ => error "Impossible case in termify_atp_abduce_candidate") |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
836 |
end |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
837 |
|
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
838 |
fun sort_top n scored_items = |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
839 |
if n <= 0 orelse null scored_items then |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
840 |
[] |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
841 |
else |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
842 |
let |
77421 | 843 |
fun split_min accum [] (_, best_item) = (best_item, List.rev accum) |
844 |
| split_min accum ((score, item) :: scored_items) (best_score, best_item) = |
|
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
845 |
if score < best_score then |
77421 | 846 |
split_min ((best_score, best_item) :: accum) scored_items (score, item) |
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
847 |
else |
77421 | 848 |
split_min ((score, item) :: accum) scored_items (best_score, best_item) |
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
849 |
|
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
850 |
val (min_item, other_scored_items) = split_min [] (tl scored_items) (hd scored_items) |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
851 |
in |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
852 |
min_item :: sort_top (n - 1) (filter_out (equal min_item o snd) other_scored_items) |
77421 | 853 |
|> distinct (op aconv) |
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
854 |
end |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
855 |
|
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
856 |
fun top_abduce_candidates max_candidates candidates = |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
857 |
let |
77421 | 858 |
(* We prefer free variables to other constructs, so that e.g. "x \<le> y" is |
77576
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
blanchet
parents:
77432
diff
changeset
|
859 |
prioritized over "x \<le> 5". *) |
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
860 |
fun score t = |
77425 | 861 |
Term.fold_aterms (fn t => fn score => score + (case t of Free _ => 1 | _ => 2)) t 0 |
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
862 |
|
77421 | 863 |
(* Equations of the form "x = ..." or "... = x" or similar are too specific |
864 |
to be useful. Quantified formulas are also filtered out. As for "True", |
|
865 |
it may seem an odd choice for abduction, but it sometimes arises in |
|
866 |
conjunction with type class constraints, which are removed by the |
|
867 |
termifier. *) |
|
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
868 |
fun maybe_score t = |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
869 |
(case t of |
77421 | 870 |
@{prop True} => NONE |
871 |
| @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ Free _ $ _) => NONE |
|
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
872 |
| @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ _ $ Free _) => NONE |
77421 | 873 |
| @{const Trueprop} $ (@{const less(nat)} $ _ $ @{const zero_class.zero(nat)}) => NONE |
874 |
| @{const Trueprop} $ (@{const less_eq(nat)} $ _ $ @{const zero_class.zero(nat)}) => NONE |
|
875 |
| @{const Trueprop} $ (@{const less(nat)} $ _ $ @{const one_class.one(nat)}) => NONE |
|
876 |
| @{const Trueprop} $ (@{const Not} $ |
|
877 |
(@{const less(nat)} $ @{const zero_class.zero(nat)} $ _)) => NONE |
|
878 |
| @{const Trueprop} $ (@{const Not} $ |
|
879 |
(@{const less_eq(nat)} $ @{const zero_class.zero(nat)} $ _)) => NONE |
|
880 |
| @{const Trueprop} $ (@{const Not} $ |
|
77425 | 881 |
(@{const less_eq(nat)} $ @{const one_class.one(nat)} $ _)) => NONE |
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
882 |
| @{const Trueprop} $ (Const (@{const_name All}, _) $ _) => NONE |
77421 | 883 |
| @{const Trueprop} $ (Const (@{const_name Ex}, _) $ _) => NONE |
77576
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
blanchet
parents:
77432
diff
changeset
|
884 |
| _ => |
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
blanchet
parents:
77432
diff
changeset
|
885 |
(* We require the presence of at least one free variable. A "missing |
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
blanchet
parents:
77432
diff
changeset
|
886 |
assumption" that does not talk about any free variable is likely |
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
blanchet
parents:
77432
diff
changeset
|
887 |
spurious. *) |
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
blanchet
parents:
77432
diff
changeset
|
888 |
if exists_subterm (fn Free _ => true | _ => false) t then SOME (score t, t) |
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
blanchet
parents:
77432
diff
changeset
|
889 |
else NONE) |
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
890 |
in |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
891 |
sort_top max_candidates (map_filter maybe_score candidates) |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
892 |
end |
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77272
diff
changeset
|
893 |
|
77602
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
894 |
fun provability_status ctxt t = |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
895 |
let |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
896 |
val res = Timeout.apply (seconds 0.1) |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
897 |
(Thm.term_of o Thm.rhs_of o Simplifier.full_rewrite ctxt) (Thm.cterm_of ctxt t) |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
898 |
in |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
899 |
if res aconv @{prop True} then SOME true |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
900 |
else if res aconv @{prop False} then SOME false |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
901 |
else NONE |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
902 |
end |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
903 |
handle Timeout.TIMEOUT _ => NONE |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
904 |
|
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
905 |
(* Put propositions that simplify to "True" first, and filter out propositions |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
906 |
that simplify to "False". *) |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
907 |
fun sort_propositions_by_provability ctxt ts = |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
908 |
let |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
909 |
val statuses_ts = map (`(provability_status ctxt)) ts |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
910 |
in |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
911 |
maps (fn (SOME true, t) => [t] | _ => []) statuses_ts @ |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
912 |
maps (fn (NONE, t) => [t] | _ => []) statuses_ts |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
913 |
end |
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
blanchet
parents:
77576
diff
changeset
|
914 |
|
31038 | 915 |
end; |