| author | blanchet | 
| Mon, 06 Oct 2014 19:19:16 +0200 | |
| changeset 58597 | 21a741e96970 | 
| parent 58500 | 430306aa03b1 | 
| child 58599 | 733b7a3277b2 | 
| 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  | 
| 
57767
 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 
blanchet 
parents: 
57765 
diff
changeset
 | 
33  | 
val simplify_bool : term -> term  | 
| 57792 | 34  | 
val var_name_of_typ : typ -> string  | 
| 57765 | 35  | 
val rename_bound_vars : term -> term  | 
| 55285 | 36  | 
val type_enc_aliases : (string * string list) list  | 
| 
45519
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
37  | 
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
 | 
38  | 
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
 | 
39  | 
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
 | 
40  | 
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
 | 
41  | 
bool -> int Symtab.table ->  | 
| 54811 | 42  | 
(string, string, (string, string atp_type) atp_term, string) atp_formula -> term  | 
| 54499 | 43  | 
|
| 57263 | 44  | 
val used_facts_in_atp_proof : Proof.context -> (string * stature) list -> string atp_proof ->  | 
45  | 
(string * stature) list  | 
|
46  | 
val used_facts_in_unsound_atp_proof : Proof.context -> (string * stature) list -> 'a atp_proof ->  | 
|
47  | 
string list option  | 
|
| 55257 | 48  | 
val atp_proof_prefers_lifting : string atp_proof -> bool  | 
| 54500 | 49  | 
val is_typed_helper_used_in_atp_proof : string atp_proof -> bool  | 
| 54772 | 50  | 
  val replace_dependencies_in_line : atp_step_name * atp_step_name list -> ('a, 'b) atp_step ->
 | 
51  | 
    ('a, 'b) atp_step
 | 
|
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
52  | 
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
 | 
53  | 
ATP_Problem_Generate.type_enc -> string Symtab.table -> (string * term) list ->  | 
| 54499 | 54  | 
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
 | 
55  | 
val repair_waldmeister_endgame : (term, 'a) atp_step list -> (term, 'a) atp_step list  | 
| 57699 | 56  | 
val infer_formulas_types : Proof.context -> term list -> term list  | 
| 57787 | 57  | 
val introduce_spass_skolems : (term, string) atp_step list -> (term, string) atp_step list  | 
| 57263 | 58  | 
val factify_atp_proof : (string * 'a) list -> term list -> term -> (term, string) atp_step list ->  | 
59  | 
(term, string) atp_step list  | 
|
| 
24425
 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 
paulson 
parents: 
24387 
diff
changeset
 | 
60  | 
end;  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
61  | 
|
| 46320 | 62  | 
structure ATP_Proof_Reconstruct : ATP_PROOF_RECONSTRUCT =  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
63  | 
struct  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
64  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43062 
diff
changeset
 | 
65  | 
open ATP_Util  | 
| 38028 | 66  | 
open ATP_Problem  | 
| 39452 | 67  | 
open ATP_Proof  | 
| 46320 | 68  | 
open ATP_Problem_Generate  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
69  | 
|
| 
45519
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
70  | 
val metisN = "metis"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
71  | 
|
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
72  | 
val full_typesN = "full_types"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
73  | 
val partial_typesN = "partial_types"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
74  | 
val no_typesN = "no_types"  | 
| 
 
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 really_full_type_enc = "mono_tags"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
77  | 
val full_type_enc = "poly_guards_query"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
78  | 
val partial_type_enc = "poly_args"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
79  | 
val no_type_enc = "erased"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
80  | 
|
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
81  | 
val full_type_encs = [full_type_enc, really_full_type_enc]  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
82  | 
val partial_type_encs = partial_type_enc :: full_type_encs  | 
| 
 
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 type_enc_aliases =  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
85  | 
[(full_typesN, full_type_encs),  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
86  | 
(partial_typesN, partial_type_encs),  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
87  | 
(no_typesN, [no_type_enc])]  | 
| 
 
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  | 
fun unalias_type_enc s =  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
90  | 
AList.lookup (op =) type_enc_aliases s |> the_default [s]  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
91  | 
|
| 54500 | 92  | 
val default_metis_lam_trans = combsN  | 
| 
45519
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
93  | 
|
| 
50670
 
eaa540986291
properly take the existential closure of skolems
 
blanchet 
parents: 
49983 
diff
changeset
 | 
94  | 
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
 | 
95  | 
| term_name' _ = ""  | 
| 
50670
 
eaa540986291
properly take the existential closure of skolems
 
blanchet 
parents: 
49983 
diff
changeset
 | 
96  | 
|
| 
 
eaa540986291
properly take the existential closure of skolems
 
blanchet 
parents: 
49983 
diff
changeset
 | 
97  | 
fun lambda' v = Term.lambda_name (term_name' v, v)  | 
| 
 
eaa540986291
properly take the existential closure of skolems
 
blanchet 
parents: 
49983 
diff
changeset
 | 
98  | 
|
| 
 
eaa540986291
properly take the existential closure of skolems
 
blanchet 
parents: 
49983 
diff
changeset
 | 
99  | 
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
 | 
100  | 
fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda' v t  | 
| 39425 | 101  | 
|
| 
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
 | 
102  | 
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
 | 
103  | 
let val ww = "'" ^ w in  | 
| 56254 | 104  | 
    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
 | 
105  | 
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
 | 
106  | 
|
| 
57767
 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 
blanchet 
parents: 
57765 
diff
changeset
 | 
107  | 
fun simplify_bool ((all as Const (@{const_name All}, _)) $ Abs (s, T, t)) =
 | 
| 
 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 
blanchet 
parents: 
57765 
diff
changeset
 | 
108  | 
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
 | 
109  | 
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
 | 
110  | 
end  | 
| 
57768
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
111  | 
  | simplify_bool (Const (@{const_name Not}, _) $ t) = s_not (simplify_bool t)
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
112  | 
  | simplify_bool (Const (@{const_name conj}, _) $ t $ u) =
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
113  | 
s_conj (simplify_bool t, simplify_bool u)  | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
114  | 
  | simplify_bool (Const (@{const_name disj}, _) $ t $ u) =
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
115  | 
s_disj (simplify_bool t, simplify_bool u)  | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
116  | 
  | simplify_bool (Const (@{const_name implies}, _) $ t $ u) =
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
117  | 
s_imp (simplify_bool t, simplify_bool u)  | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
118  | 
  | simplify_bool ((t as Const (@{const_name HOL.eq}, _)) $ u $ v) =
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
119  | 
(case (u, v) of  | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
120  | 
      (Const (@{const_name True}, _), _) => v
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
121  | 
    | (u, Const (@{const_name True}, _)) => u
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
122  | 
    | (Const (@{const_name False}, _), v) => s_not v
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
123  | 
    | (u, Const (@{const_name False}, _)) => s_not u
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
124  | 
    | _ => if u aconv v then @{const True} 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
 | 
125  | 
| 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
 | 
126  | 
| 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
 | 
127  | 
| simplify_bool t = t  | 
| 
 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 
blanchet 
parents: 
57765 
diff
changeset
 | 
128  | 
|
| 57765 | 129  | 
fun single_letter upper s =  | 
130  | 
let val s' = if String.isPrefix "o" s orelse String.isPrefix "O" s then "z" else s in  | 
|
131  | 
String.extract (Name.desymbolize (SOME upper) (Long_Name.base_name s'), 0, SOME 1)  | 
|
132  | 
end  | 
|
133  | 
||
134  | 
fun var_name_of_typ (Type (@{type_name fun}, [_, T])) =
 | 
|
135  | 
if body_type T = HOLogic.boolT then "p" else "f"  | 
|
136  | 
  | var_name_of_typ (Type (@{type_name set}, [T])) = single_letter true (var_name_of_typ T)
 | 
|
137  | 
| var_name_of_typ (Type (s, Ts)) =  | 
|
138  | 
if String.isSuffix "list" s then var_name_of_typ (the_single Ts) ^ "s"  | 
|
139  | 
else single_letter false (Long_Name.base_name s)  | 
|
140  | 
| var_name_of_typ (TFree (s, _)) = single_letter false (perhaps (try (unprefix "'")) s)  | 
|
141  | 
| var_name_of_typ (TVar ((s, _), T)) = var_name_of_typ (TFree (s, T))  | 
|
142  | 
||
143  | 
fun rename_bound_vars (t $ u) = rename_bound_vars t $ rename_bound_vars u  | 
|
144  | 
| rename_bound_vars (Abs (_, T, t)) = Abs (var_name_of_typ T, T, rename_bound_vars t)  | 
|
145  | 
| rename_bound_vars t = t  | 
|
146  | 
||
| 54820 | 147  | 
exception ATP_CLASS of string list  | 
| 54818 | 148  | 
exception ATP_TYPE of string atp_type list  | 
| 54811 | 149  | 
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
 | 
150  | 
exception ATP_FORMULA of  | 
| 54811 | 151  | 
(string, string, (string, string atp_type) atp_term, string) atp_formula list  | 
| 37991 | 152  | 
exception SAME of unit  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
153  | 
|
| 54820 | 154  | 
fun class_of_atp_class cls =  | 
155  | 
(case unprefix_and_unascii class_prefix cls of  | 
|
156  | 
SOME s => s  | 
|
157  | 
| NONE => raise ATP_CLASS [cls])  | 
|
158  | 
||
| 
57697
 
44341963ade3
correctly translate THF functions from terms to types
 
blanchet 
parents: 
57635 
diff
changeset
 | 
159  | 
fun atp_type_of_atp_term (ATerm ((s, _), us)) =  | 
| 
 
44341963ade3
correctly translate THF functions from terms to types
 
blanchet 
parents: 
57635 
diff
changeset
 | 
160  | 
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
 | 
161  | 
if s = tptp_fun_type then  | 
| 
 
44341963ade3
correctly translate THF functions from terms to types
 
blanchet 
parents: 
57635 
diff
changeset
 | 
162  | 
(case tys of  | 
| 
 
44341963ade3
correctly translate THF functions from terms to types
 
blanchet 
parents: 
57635 
diff
changeset
 | 
163  | 
[ty1, ty2] => AFun (ty1, ty2)  | 
| 
 
44341963ade3
correctly translate THF functions from terms to types
 
blanchet 
parents: 
57635 
diff
changeset
 | 
164  | 
| _ => raise ATP_TYPE tys)  | 
| 
 
44341963ade3
correctly translate THF functions from terms to types
 
blanchet 
parents: 
57635 
diff
changeset
 | 
165  | 
else  | 
| 
 
44341963ade3
correctly translate THF functions from terms to types
 
blanchet 
parents: 
57635 
diff
changeset
 | 
166  | 
AType ((s, []), tys)  | 
| 37991 | 167  | 
end  | 
| 54818 | 168  | 
|
| 58477 | 169  | 
(* Type variables are given the basic sort "HOL.type". Some will later be constrained by information  | 
170  | 
from type literals, or by type inference. *)  | 
|
171  | 
fun typ_of_atp_type ctxt (ty as AType ((a, clss), tys)) =  | 
|
172  | 
let val Ts = map (typ_of_atp_type ctxt) tys in  | 
|
173  | 
(case unprefix_and_unascii native_type_prefix a of  | 
|
174  | 
SOME b => typ_of_atp_type ctxt (atp_type_of_atp_term (unmangled_type b))  | 
|
175  | 
| NONE =>  | 
|
176  | 
(case unprefix_and_unascii type_const_prefix a of  | 
|
177  | 
SOME b => Type (invert_const b, Ts)  | 
|
178  | 
| NONE =>  | 
|
179  | 
if not (null tys) then  | 
|
180  | 
raise ATP_TYPE [ty] (* only "tconst"s have type arguments *)  | 
|
181  | 
else  | 
|
182  | 
(case unprefix_and_unascii tfree_prefix a of  | 
|
183  | 
SOME b => make_tfree ctxt b  | 
|
184  | 
| NONE =>  | 
|
185  | 
(* The term could be an Isabelle variable or a variable from the ATP, say "X1" or "_5018".  | 
|
186  | 
Sometimes variables from the ATP are indistinguishable from Isabelle variables, which  | 
|
187  | 
forces us to use a type parameter in all cases. *)  | 
|
188  | 
              Type_Infer.param 0 ("'" ^ perhaps (unprefix_and_unascii tvar_prefix) a,
 | 
|
189  | 
                (if null clss then @{sort type} else map class_of_atp_class clss)))))
 | 
|
190  | 
end  | 
|
191  | 
| typ_of_atp_type ctxt (AFun (ty1, ty2)) = typ_of_atp_type ctxt ty1 --> typ_of_atp_type ctxt ty2  | 
|
192  | 
||
| 54818 | 193  | 
fun typ_of_atp_term ctxt = typ_of_atp_type ctxt o atp_type_of_atp_term  | 
194  | 
||
| 54789 | 195  | 
(* 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
 | 
196  | 
fun type_constraint_of_term ctxt (u as ATerm ((a, _), us)) =  | 
| 54818 | 197  | 
(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
 | 
198  | 
(SOME b, [T]) => (b, T)  | 
| 54789 | 199  | 
| _ => raise ATP_TERM [u])  | 
| 38014 | 200  | 
|
| 
43907
 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 
blanchet 
parents: 
43863 
diff
changeset
 | 
201  | 
(* 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
 | 
202  | 
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
 | 
203  | 
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
 | 
204  | 
| 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
 | 
205  | 
| add_type_constraint _ _ = I  | 
| 38014 | 206  | 
|
| 
57789
 
a73255cffb5b
don't rename variables which will be renamed later anyway
 
blanchet 
parents: 
57788 
diff
changeset
 | 
207  | 
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
 | 
208  | 
(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
 | 
209  | 
SOME s' => s'  | 
| 
 
a73255cffb5b
don't rename variables which will be renamed later anyway
 
blanchet 
parents: 
57788 
diff
changeset
 | 
210  | 
| NONE => s)  | 
| 
55185
 
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
 
blanchet 
parents: 
54822 
diff
changeset
 | 
211  | 
|
| 54789 | 212  | 
(* The number of type arguments of a constant, zero if it's monomorphic. For (instances of) Skolem  | 
213  | 
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
 | 
214  | 
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
 | 
215  | 
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
 | 
216  | 
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
 | 
217  | 
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
 | 
218  | 
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
 | 
219  | 
else  | 
| 
 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 
blanchet 
parents: 
43863 
diff
changeset
 | 
220  | 
(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
 | 
221  | 
|
| 56254 | 222  | 
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
 | 
223  | 
|
| 54772 | 224  | 
val spass_skolem_prefix = "sk" (* "skc" or "skf" *)  | 
225  | 
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
 | 
226  | 
|
| 57257 | 227  | 
fun var_index_of_textual textual = if textual then 0 else 1  | 
228  | 
||
229  | 
fun quantify_over_var textual quant_of var_s var_T t =  | 
|
230  | 
let  | 
|
| 57790 | 231  | 
val vars =  | 
232  | 
((var_s, var_index_of_textual textual), var_T) ::  | 
|
233  | 
filter (fn ((s, _), _) => s = var_s) (Term.add_vars t [])  | 
|
| 57257 | 234  | 
val normTs = vars |> AList.group (op =) |> map (apsnd hd)  | 
235  | 
fun norm_var_types (Var (x, T)) =  | 
|
236  | 
Var (x, the_default T (AList.lookup (op =) normTs x))  | 
|
237  | 
| norm_var_types t = t  | 
|
238  | 
in t |> map_aterms norm_var_types |> fold_rev quant_of (map Var normTs) end  | 
|
239  | 
||
| 57790 | 240  | 
(* This assumes that distinct names are mapped to distinct names by "Variable.variant_frees". This  | 
241  | 
does not hold in general but should hold for ATP-generated Skolem function names, since these end  | 
|
242  | 
with a digit and "variant_frees" appends letters. *)  | 
|
243  | 
fun fresh_up ctxt s = fst (hd (Variable.variant_frees ctxt [] [(s, ())]))  | 
|
244  | 
||
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
245  | 
(* 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
 | 
246  | 
are typed. The code is similar to "term_of_atp_fo". *)  | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
247  | 
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
 | 
248  | 
let  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
249  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
57768
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
250  | 
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
 | 
251  | 
|
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
252  | 
fun do_term opt_T u =  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
253  | 
(case u of  | 
| 
57788
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
254  | 
AAbs (((var, ty), term), []) =>  | 
| 57257 | 255  | 
let  | 
256  | 
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
 | 
257  | 
val var_name = repair_var_name var  | 
| 57257 | 258  | 
val tm = do_term NONE term  | 
| 
57768
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
259  | 
in quantify_over_var true lambda' var_name typ tm end  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
260  | 
| ATerm ((s, tys), us) =>  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
261  | 
if s = ""  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
262  | 
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
 | 
263  | 
\contains unparsable material."  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
264  | 
else if s = tptp_equal then  | 
| 
57768
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
265  | 
          list_comb (Const (@{const_name HOL.eq}, Type_Infer.anyT @{sort type}),
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
266  | 
map (do_term NONE) us)  | 
| 57256 | 267  | 
else if not (null us) then  | 
268  | 
let  | 
|
| 57820 | 269  | 
val args = map (do_term NONE) us  | 
| 57256 | 270  | 
val opt_T' = SOME (map slack_fastype_of args ---> the_default dummyT opt_T)  | 
271  | 
val func = do_term opt_T' (ATerm ((s, tys), []))  | 
|
272  | 
in foldl1 (op $) (func :: args) end  | 
|
273  | 
else if s = tptp_or then HOLogic.disj  | 
|
274  | 
else if s = tptp_and then HOLogic.conj  | 
|
275  | 
else if s = tptp_implies then HOLogic.imp  | 
|
276  | 
else if s = tptp_iff orelse s = tptp_equal then HOLogic.eq_const dummyT  | 
|
| 
57768
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
277  | 
        else if s = tptp_not_iff orelse s = tptp_not_equal then @{term "%P Q. Q ~= P"}
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
278  | 
        else if s = tptp_if then @{term "%P Q. Q --> P"}
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
279  | 
        else if s = tptp_not_and then @{term "%P Q. ~ (P & Q)"}
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
280  | 
        else if s = tptp_not_or then @{term "%P Q. ~ (P | Q)"}
 | 
| 57256 | 281  | 
else if s = tptp_not then HOLogic.Not  | 
282  | 
else if s = tptp_ho_forall then HOLogic.all_const dummyT  | 
|
283  | 
else if s = tptp_ho_exists then HOLogic.exists_const dummyT  | 
|
| 57709 | 284  | 
else if s = tptp_hilbert_choice then HOLogic.choice_const dummyT  | 
285  | 
        else if s = tptp_hilbert_the then @{term "The"}
 | 
|
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
286  | 
else  | 
| 57256 | 287  | 
(case unprefix_and_unascii const_prefix s of  | 
288  | 
SOME s' =>  | 
|
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
289  | 
let  | 
| 
57547
 
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
 
blanchet 
parents: 
57267 
diff
changeset
 | 
290  | 
val ((s', _), mangled_us) = s' |> unmangled_const |>> `invert_const  | 
| 57256 | 291  | 
val num_ty_args = length us - the_default 0 (Symtab.lookup sym_tab s)  | 
292  | 
val (type_us, term_us) = chop num_ty_args us |>> append mangled_us  | 
|
293  | 
val term_ts = map (do_term NONE) term_us  | 
|
294  | 
val Ts = map (typ_of_atp_type ctxt) tys @ map (typ_of_atp_term ctxt) type_us  | 
|
295  | 
val T =  | 
|
296  | 
(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
 | 
297  | 
try (Sign.const_instance thy) (s', Ts)  | 
| 57256 | 298  | 
else  | 
299  | 
NONE)  | 
|
300  | 
|> (fn SOME T => T  | 
|
301  | 
| NONE =>  | 
|
302  | 
map slack_fastype_of term_ts --->  | 
|
303  | 
                       the_default (Type_Infer.anyT @{sort type}) opt_T)
 | 
|
304  | 
val t = Const (unproxify_const s', T)  | 
|
305  | 
in list_comb (t, term_ts) end  | 
|
306  | 
| NONE => (* a free or schematic variable *)  | 
|
307  | 
let  | 
|
308  | 
val ts = map (do_term NONE) us  | 
|
309  | 
val T =  | 
|
| 
58500
 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 
blanchet 
parents: 
58484 
diff
changeset
 | 
310  | 
(case tys of  | 
| 
 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 
blanchet 
parents: 
58484 
diff
changeset
 | 
311  | 
[ty] => typ_of_atp_type ctxt ty  | 
| 
 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 
blanchet 
parents: 
58484 
diff
changeset
 | 
312  | 
| _ =>  | 
| 57256 | 313  | 
map slack_fastype_of ts --->  | 
| 
58500
 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 
blanchet 
parents: 
58484 
diff
changeset
 | 
314  | 
(case opt_T of  | 
| 
 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 
blanchet 
parents: 
58484 
diff
changeset
 | 
315  | 
SOME T => T  | 
| 
 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 
blanchet 
parents: 
58484 
diff
changeset
 | 
316  | 
                  | NONE => Type_Infer.anyT @{sort type}))
 | 
| 57256 | 317  | 
val t =  | 
318  | 
(case unprefix_and_unascii fixed_var_prefix s of  | 
|
319  | 
SOME s => Free (s, T)  | 
|
320  | 
| NONE =>  | 
|
| 57790 | 321  | 
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
 | 
322  | 
else Var ((repair_var_name s, var_index), T))  | 
| 57256 | 323  | 
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
 | 
324  | 
in do_term end  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
325  | 
|
| 56254 | 326  | 
(* First-order translation. No types are known for variables. "Type_Infer.anyT @{sort type}"
 | 
327  | 
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
 | 
328  | 
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
 | 
329  | 
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
 | 
330  | 
val thy = Proof_Context.theory_of ctxt  | 
| 54789 | 331  | 
(* For Metis, we use 1 rather than 0 because variable references in clauses may otherwise  | 
332  | 
conflict with variable constraints in the goal. At least, type inference often fails  | 
|
333  | 
otherwise. See also "axiom_inference" in "Metis_Reconstruct". *)  | 
|
| 57257 | 334  | 
val var_index = var_index_of_textual textual  | 
| 57199 | 335  | 
|
| 
43131
 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 
blanchet 
parents: 
43130 
diff
changeset
 | 
336  | 
fun do_term extra_ts opt_T u =  | 
| 54789 | 337  | 
(case u of  | 
| 54818 | 338  | 
ATerm ((s, tys), us) =>  | 
| 57199 | 339  | 
if s = "" then  | 
340  | 
error "Isar proof reconstruction failed because the ATP proof contains unparsable \  | 
|
341  | 
\material."  | 
|
| 51878 | 342  | 
else if String.isPrefix native_type_prefix s then  | 
| 58477 | 343  | 
          @{const True} (* 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
 | 
344  | 
else if s = tptp_equal then  | 
| 
57768
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
345  | 
          list_comb (Const (@{const_name HOL.eq}, Type_Infer.anyT @{sort type}),
 | 
| 
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
346  | 
map (do_term [] NONE) us)  | 
| 54789 | 347  | 
else  | 
348  | 
(case unprefix_and_unascii const_prefix s of  | 
|
349  | 
SOME s' =>  | 
|
| 54818 | 350  | 
let val ((s', s''), mangled_us) = s' |> unmangled_const |>> `invert_const in  | 
| 54789 | 351  | 
if s' = type_tag_name then  | 
352  | 
(case mangled_us @ us of  | 
|
| 54818 | 353  | 
[typ_u, term_u] => do_term extra_ts (SOME (typ_of_atp_term ctxt typ_u)) term_u  | 
| 54789 | 354  | 
| _ => raise ATP_TERM us)  | 
355  | 
else if s' = predicator_name then  | 
|
356  | 
                do_term [] (SOME @{typ bool}) (hd us)
 | 
|
357  | 
else if s' = app_op_name then  | 
|
358  | 
let val extra_t = do_term [] NONE (List.last us) in  | 
|
359  | 
do_term (extra_t :: extra_ts)  | 
|
| 57199 | 360  | 
(case opt_T of SOME T => SOME (slack_fastype_of extra_t --> T) | NONE => NONE)  | 
361  | 
(nth us (length us - 2))  | 
|
| 54789 | 362  | 
end  | 
363  | 
else if s' = type_guard_name then  | 
|
364  | 
                @{const True} (* ignore type predicates *)
 | 
|
365  | 
else  | 
|
366  | 
let  | 
|
367  | 
val new_skolem = String.isPrefix new_skolem_const_prefix s''  | 
|
| 54818 | 368  | 
val num_ty_args = length us - the_default 0 (Symtab.lookup sym_tab s)  | 
369  | 
val (type_us, term_us) = chop num_ty_args us |>> append mangled_us  | 
|
| 54789 | 370  | 
val term_ts = map (do_term [] NONE) term_us  | 
| 54818 | 371  | 
|
372  | 
val Ts = map (typ_of_atp_type ctxt) tys @ map (typ_of_atp_term ctxt) type_us  | 
|
| 54789 | 373  | 
val T =  | 
| 54818 | 374  | 
(if not (null Ts) andalso robust_const_num_type_args thy s' = length Ts then  | 
| 57199 | 375  | 
if new_skolem then SOME (Type_Infer.paramify_vars (tl Ts ---> hd Ts))  | 
376  | 
else if textual then try (Sign.const_instance thy) (s', Ts)  | 
|
377  | 
else NONE  | 
|
| 54789 | 378  | 
else  | 
379  | 
NONE)  | 
|
380  | 
|> (fn SOME T => T  | 
|
381  | 
| NONE =>  | 
|
| 56254 | 382  | 
map slack_fastype_of term_ts --->  | 
| 57199 | 383  | 
                           the_default (Type_Infer.anyT @{sort type}) opt_T)
 | 
| 54789 | 384  | 
val t =  | 
385  | 
if new_skolem then Var ((new_skolem_var_name_of_const s'', var_index), T)  | 
|
386  | 
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
 | 
387  | 
in  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
388  | 
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
 | 
389  | 
end  | 
| 54789 | 390  | 
end  | 
391  | 
| NONE => (* a free or schematic variable *)  | 
|
392  | 
let  | 
|
393  | 
val term_ts =  | 
|
394  | 
map (do_term [] NONE) us  | 
|
395  | 
(* SPASS (3.8ds) and Vampire (2.6) pass arguments to Skolem functions in reverse  | 
|
| 57699 | 396  | 
order, which is incompatible with "metis"'s new skolemizer. *)  | 
| 54789 | 397  | 
|> exists (fn pre => String.isPrefix pre s)  | 
398  | 
[spass_skolem_prefix, vampire_skolem_prefix] ? rev  | 
|
399  | 
val ts = term_ts @ extra_ts  | 
|
400  | 
val T =  | 
|
| 
58500
 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 
blanchet 
parents: 
58484 
diff
changeset
 | 
401  | 
(case tys of  | 
| 
 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 
blanchet 
parents: 
58484 
diff
changeset
 | 
402  | 
[ty] => typ_of_atp_type ctxt ty  | 
| 
 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 
blanchet 
parents: 
58484 
diff
changeset
 | 
403  | 
| _ =>  | 
| 
 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 
blanchet 
parents: 
58484 
diff
changeset
 | 
404  | 
(case opt_T of  | 
| 
 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 
blanchet 
parents: 
58484 
diff
changeset
 | 
405  | 
SOME T => map slack_fastype_of term_ts ---> T  | 
| 
 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 
blanchet 
parents: 
58484 
diff
changeset
 | 
406  | 
                  | NONE => map slack_fastype_of ts ---> Type_Infer.anyT @{sort type}))
 | 
| 54789 | 407  | 
val t =  | 
408  | 
(case unprefix_and_unascii fixed_var_prefix s of  | 
|
409  | 
SOME s => Free (s, T)  | 
|
| 
36909
 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 
blanchet 
parents: 
36607 
diff
changeset
 | 
410  | 
| NONE =>  | 
| 
55185
 
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
 
blanchet 
parents: 
54822 
diff
changeset
 | 
411  | 
if textual andalso not (is_tptp_variable s) then  | 
| 57790 | 412  | 
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
 | 
413  | 
else  | 
| 
57789
 
a73255cffb5b
don't rename variables which will be renamed later anyway
 
blanchet 
parents: 
57788 
diff
changeset
 | 
414  | 
Var ((repair_var_name s, var_index), T))  | 
| 54789 | 415  | 
in list_comb (t, ts) end))  | 
| 43093 | 416  | 
in do_term [] end  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
417  | 
|
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
418  | 
fun term_of_atp ctxt (ATP_Problem.THF _) type_enc =  | 
| 
57768
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
419  | 
if ATP_Problem_Generate.is_type_enc_higher_order type_enc then K (term_of_atp_ho ctxt)  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
420  | 
else error "Unsupported Isar reconstruction."  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
421  | 
| term_of_atp ctxt _ type_enc =  | 
| 
57768
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
422  | 
if not (ATP_Problem_Generate.is_type_enc_higher_order type_enc) then term_of_atp_fo ctxt  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
423  | 
else error "Unsupported Isar reconstruction."  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
424  | 
|
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
425  | 
fun term_of_atom ctxt format type_enc textual sym_tab pos (u as ATerm ((s, _), _)) =  | 
| 38014 | 426  | 
if String.isPrefix class_prefix s then  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
51878 
diff
changeset
 | 
427  | 
add_type_constraint pos (type_constraint_of_term ctxt u)  | 
| 38014 | 428  | 
    #> pair @{const True}
 | 
429  | 
else  | 
|
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
430  | 
    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
 | 
431  | 
|
| 37991 | 432  | 
(* Update schematic type variables with detected sort constraints. It's not  | 
| 42563 | 433  | 
totally clear whether this code is necessary. *)  | 
| 38014 | 434  | 
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
 | 
435  | 
let  | 
| 37991 | 436  | 
fun do_type (Type (a, Ts)) = Type (a, map do_type Ts)  | 
437  | 
| do_type (TVar (xi, s)) =  | 
|
438  | 
TVar (xi, the_default s (Vartab.lookup tvar_tab xi))  | 
|
439  | 
| do_type (TFree z) = TFree z  | 
|
440  | 
fun do_term (Const (a, T)) = Const (a, do_type T)  | 
|
441  | 
| do_term (Free (a, T)) = Free (a, do_type T)  | 
|
442  | 
| do_term (Var (xi, T)) = Var (xi, do_type T)  | 
|
443  | 
| do_term (t as Bound _) = t  | 
|
444  | 
| do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t)  | 
|
445  | 
| do_term (t1 $ t2) = do_term t1 $ do_term t2  | 
|
446  | 
in t |> not (Vartab.is_empty tvar_tab) ? do_term end  | 
|
447  | 
||
| 54811 | 448  | 
(* Interpret an ATP formula as a HOL term, extracting sort constraints as they appear in the  | 
449  | 
formula. *)  | 
|
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
450  | 
fun prop_of_atp ctxt format type_enc textual sym_tab phi =  | 
| 38014 | 451  | 
let  | 
452  | 
fun do_formula pos phi =  | 
|
| 54789 | 453  | 
(case phi of  | 
| 38014 | 454  | 
AQuant (_, [], phi) => do_formula pos phi  | 
| 42526 | 455  | 
| AQuant (q, (s, _) :: xs, phi') =>  | 
| 38014 | 456  | 
do_formula pos (AQuant (q, xs, phi'))  | 
| 42526 | 457  | 
(* FIXME: TFF *)  | 
| 57257 | 458  | 
#>> 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
 | 
459  | 
(repair_var_name s) dummyT  | 
| 38014 | 460  | 
| AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not  | 
| 37991 | 461  | 
| AConn (c, [phi1, phi2]) =>  | 
| 38014 | 462  | 
do_formula (pos |> c = AImplies ? not) phi1  | 
463  | 
##>> do_formula pos phi2  | 
|
464  | 
#>> (case c of  | 
|
| 54811 | 465  | 
AAnd => s_conj  | 
466  | 
| AOr => s_disj  | 
|
467  | 
| AImplies => s_imp  | 
|
468  | 
| AIff => s_iff  | 
|
469  | 
| ANot => raise Fail "impossible connective")  | 
|
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
470  | 
| AAtom tm => term_of_atom ctxt format type_enc textual sym_tab pos tm  | 
| 54789 | 471  | 
| _ => raise ATP_FORMULA [phi])  | 
472  | 
in  | 
|
473  | 
repair_tvar_sorts (do_formula true phi Vartab.empty)  | 
|
474  | 
end  | 
|
| 37991 | 475  | 
|
| 54500 | 476  | 
val unprefix_fact_number = space_implode "_" o tl o space_explode "_"  | 
477  | 
||
| 57263 | 478  | 
fun resolve_fact facts s =  | 
| 54789 | 479  | 
(case try (unprefix fact_prefix) s of  | 
| 54500 | 480  | 
SOME s' =>  | 
481  | 
let val s' = s' |> unprefix_fact_number |> unascii_of in  | 
|
| 57263 | 482  | 
AList.lookup (op =) facts s' |> Option.map (pair s')  | 
| 54500 | 483  | 
end  | 
| 54789 | 484  | 
| NONE => NONE)  | 
| 54506 | 485  | 
|
| 57263 | 486  | 
fun resolve_conjecture s =  | 
| 54789 | 487  | 
(case try (unprefix conjecture_prefix) s of  | 
| 54500 | 488  | 
SOME s' => Int.fromString s'  | 
| 54789 | 489  | 
| NONE => NONE)  | 
| 54500 | 490  | 
|
| 57263 | 491  | 
fun resolve_facts facts = map_filter (resolve_fact facts)  | 
492  | 
val resolve_conjectures = map_filter resolve_conjecture  | 
|
| 54500 | 493  | 
|
494  | 
fun is_axiom_used_in_proof pred =  | 
|
495  | 
exists (fn ((_, ss), _, _, _, []) => exists pred ss | _ => false)  | 
|
496  | 
||
497  | 
val leo2_extcnf_equal_neg_rule = "extcnf_equal_neg"  | 
|
498  | 
||
| 57263 | 499  | 
fun add_fact ctxt facts ((_, ss), _, _, rule, deps) =  | 
| 
57707
 
0242e9578828
imported patch satallax_proof_support_Sledgehammer
 
fleury 
parents: 
57703 
diff
changeset
 | 
500  | 
(if member (op =) [agsyhol_core_rule, leo2_extcnf_equal_neg_rule] rule then  | 
| 
56104
 
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
 
blanchet 
parents: 
55285 
diff
changeset
 | 
501  | 
insert (op =) (short_thm_name ctxt ext, (Global, General))  | 
| 54500 | 502  | 
else  | 
503  | 
I)  | 
|
| 57263 | 504  | 
#> (if null deps then union (op =) (resolve_facts facts ss) else I)  | 
| 54500 | 505  | 
|
| 57263 | 506  | 
fun used_facts_in_atp_proof ctxt facts atp_proof =  | 
507  | 
if null atp_proof then facts else fold (add_fact ctxt facts) atp_proof []  | 
|
| 54500 | 508  | 
|
509  | 
fun used_facts_in_unsound_atp_proof _ _ [] = NONE  | 
|
| 57263 | 510  | 
| used_facts_in_unsound_atp_proof ctxt facts atp_proof =  | 
511  | 
let val used_facts = used_facts_in_atp_proof ctxt facts atp_proof in  | 
|
| 54500 | 512  | 
if forall (fn (_, (sc, _)) => sc = Global) used_facts andalso  | 
| 57263 | 513  | 
not (is_axiom_used_in_proof (is_some o resolve_conjecture) atp_proof) then  | 
| 54500 | 514  | 
SOME (map fst used_facts)  | 
515  | 
else  | 
|
516  | 
NONE  | 
|
517  | 
end  | 
|
518  | 
||
519  | 
val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix  | 
|
520  | 
||
521  | 
(* overapproximation (good enough) *)  | 
|
522  | 
fun is_lam_lifted s =  | 
|
523  | 
String.isPrefix fact_prefix s andalso  | 
|
524  | 
String.isSubstring ascii_of_lam_fact_prefix s  | 
|
525  | 
||
526  | 
val is_combinator_def = String.isPrefix (helper_prefix ^ combinator_prefix)  | 
|
527  | 
||
| 55257 | 528  | 
fun atp_proof_prefers_lifting atp_proof =  | 
529  | 
(is_axiom_used_in_proof is_combinator_def atp_proof,  | 
|
530  | 
is_axiom_used_in_proof is_lam_lifted atp_proof) = (false, true)  | 
|
| 54500 | 531  | 
|
532  | 
val is_typed_helper_name =  | 
|
533  | 
String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix  | 
|
534  | 
||
535  | 
fun is_typed_helper_used_in_atp_proof atp_proof =  | 
|
536  | 
is_axiom_used_in_proof is_typed_helper_name atp_proof  | 
|
537  | 
||
| 54772 | 538  | 
fun replace_one_dependency (old, new) dep = if is_same_atp_step dep old then new else [dep]  | 
539  | 
fun replace_dependencies_in_line old_new (name, role, t, rule, deps) =  | 
|
540  | 
(name, role, t, rule, fold (union (op =) o replace_one_dependency old_new) deps [])  | 
|
541  | 
||
| 54499 | 542  | 
fun repair_name "$true" = "c_True"  | 
543  | 
| repair_name "$false" = "c_False"  | 
|
544  | 
| repair_name "$$e" = tptp_equal (* seen in Vampire proofs *)  | 
|
545  | 
| repair_name s =  | 
|
546  | 
if is_tptp_equal s orelse  | 
|
547  | 
(* seen in Vampire proofs *)  | 
|
548  | 
(String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then  | 
|
549  | 
tptp_equal  | 
|
550  | 
else  | 
|
551  | 
s  | 
|
552  | 
||
| 57699 | 553  | 
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
 | 
554  | 
|
| 57699 | 555  | 
fun infer_formulas_types ctxt =  | 
| 
57635
 
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
 
blanchet 
parents: 
57547 
diff
changeset
 | 
556  | 
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
 | 
557  | 
#> Syntax.check_terms (Proof_Context.set_mode Proof_Context.mode_schematic ctxt)  | 
| 
 
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
 
blanchet 
parents: 
57547 
diff
changeset
 | 
558  | 
#> map (set_var_index 0)  | 
| 54499 | 559  | 
|
560  | 
val combinator_table =  | 
|
561  | 
  [(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def [abs_def]}),
 | 
|
562  | 
   (@{const_name Meson.COMBK}, @{thm Meson.COMBK_def [abs_def]}),
 | 
|
563  | 
   (@{const_name Meson.COMBB}, @{thm Meson.COMBB_def [abs_def]}),
 | 
|
564  | 
   (@{const_name Meson.COMBC}, @{thm Meson.COMBC_def [abs_def]}),
 | 
|
565  | 
   (@{const_name Meson.COMBS}, @{thm Meson.COMBS_def [abs_def]})]
 | 
|
566  | 
||
567  | 
fun uncombine_term thy =  | 
|
568  | 
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
 | 
569  | 
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
 | 
570  | 
| 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
 | 
571  | 
| uncomb (t as Const (x as (s, _))) =  | 
| 54499 | 572  | 
(case AList.lookup (op =) combinator_table s of  | 
| 54756 | 573  | 
SOME thm => thm |> prop_of |> specialize_type thy x |> Logic.dest_equals |> snd  | 
574  | 
| 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
 | 
575  | 
| 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
 | 
576  | 
in uncomb end  | 
| 54499 | 577  | 
|
| 
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
 | 
578  | 
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
 | 
579  | 
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
 | 
580  | 
(* 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
 | 
581  | 
(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
 | 
582  | 
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
 | 
583  | 
| 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
 | 
584  | 
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
 | 
585  | 
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
 | 
586  | 
| 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
 | 
587  | 
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
 | 
588  | 
map_aterms (unlift_aterm lifted)  | 
| 54499 | 589  | 
|
| 57256 | 590  | 
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
 | 
591  | 
| 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
 | 
592  | 
let  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
593  | 
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
 | 
594  | 
val t = u  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
595  | 
|> 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
 | 
596  | 
|> unlift_term lifted  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
597  | 
|> uncombine_term thy  | 
| 
57768
 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 
blanchet 
parents: 
57767 
diff
changeset
 | 
598  | 
|> simplify_bool  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
599  | 
in  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
600  | 
SOME (name, role, t, rule, deps)  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
57199 
diff
changeset
 | 
601  | 
end  | 
| 54499 | 602  | 
|
603  | 
val waldmeister_conjecture_num = "1.0.0.0"  | 
|
604  | 
||
| 54756 | 605  | 
fun repair_waldmeister_endgame proof =  | 
| 54499 | 606  | 
let  | 
| 56854 | 607  | 
    fun repair_tail (name, _, @{const Trueprop} $ t, rule, deps) =
 | 
608  | 
      (name, Negated_Conjecture, @{const Trueprop} $ s_not t, rule, deps)
 | 
|
| 54756 | 609  | 
fun repair_body [] = []  | 
610  | 
| repair_body ((line as ((num, _), _, _, _, _)) :: lines) =  | 
|
611  | 
if num = waldmeister_conjecture_num then map repair_tail (line :: lines)  | 
|
612  | 
else line :: repair_body lines  | 
|
613  | 
in  | 
|
614  | 
repair_body proof  | 
|
615  | 
end  | 
|
| 54499 | 616  | 
|
| 
57635
 
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
 
blanchet 
parents: 
57547 
diff
changeset
 | 
617  | 
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
 | 
618  | 
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
 | 
619  | 
|
| 
57258
 
67d85a8aa6cc
Moving the remote prefix deleting on Sledgehammer's side
 
fleury 
parents: 
57257 
diff
changeset
 | 
620  | 
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
 | 
621  | 
nasty_atp_proof pool  | 
| 54499 | 622  | 
#> 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
 | 
623  | 
#> map_filter (termify_line ctxt format type_enc lifted sym_tab)  | 
| 57699 | 624  | 
#> 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
 | 
625  | 
#> local_prover = waldmeisterN ? repair_waldmeister_endgame  | 
| 54499 | 626  | 
|
| 58091 | 627  | 
fun unskolemize_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
 | 
628  | 
let  | 
| 58597 | 629  | 
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
 | 
630  | 
|
| 58091 | 631  | 
fun find_argless_skolem (Free _ $ Var _) = NONE  | 
| 58597 | 632  | 
| find_argless_skolem (Free (x as (s, _))) = if is_skolem_name s then SOME x else NONE  | 
| 58091 | 633  | 
| find_argless_skolem (t $ u) =  | 
634  | 
(case find_argless_skolem t of NONE => find_argless_skolem u | sk => sk)  | 
|
635  | 
| find_argless_skolem (Abs (_, _, t)) = find_argless_skolem t  | 
|
636  | 
| 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
 | 
637  | 
|
| 58597 | 638  | 
fun find_skolem_arg (Free (s, _) $ Var v) = if is_skolem_name s then SOME v else NONE  | 
| 58091 | 639  | 
| find_skolem_arg (t $ u) = (case find_skolem_arg t of NONE => find_skolem_arg u | v => v)  | 
640  | 
| find_skolem_arg (Abs (_, _, t)) = find_skolem_arg t  | 
|
641  | 
| find_skolem_arg _ = NONE  | 
|
642  | 
||
| 58597 | 643  | 
fun kill_skolem_arg (t as Free (s, T) $ Var _) =  | 
644  | 
if is_skolem_name s then Free (s, range_type T) else t  | 
|
645  | 
| kill_skolem_arg (t $ u) = kill_skolem_arg t $ kill_skolem_arg u  | 
|
646  | 
| kill_skolem_arg (Abs (s, T, t)) = Abs (s, T, kill_skolem_arg t)  | 
|
647  | 
| kill_skolem_arg t = t  | 
|
648  | 
||
| 58091 | 649  | 
fun find_var (Var v) = SOME v  | 
650  | 
| find_var (t $ u) = (case find_var t of NONE => find_var u | v => v)  | 
|
651  | 
| find_var (Abs (_, _, t)) = find_var t  | 
|
652  | 
| find_var _ = NONE  | 
|
653  | 
||
654  | 
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
 | 
655  | 
|
| 58091 | 656  | 
fun unskolem t =  | 
657  | 
(case find_argless_skolem t of  | 
|
658  | 
SOME (x as (s, T)) =>  | 
|
659  | 
HOLogic.exists_const T $ Abs (s, T, unskolem (safe_abstract_over (Free x, t)))  | 
|
660  | 
| NONE =>  | 
|
| 58597 | 661  | 
(case find_skolem_arg t of  | 
| 58091 | 662  | 
SOME (v as ((s, _), T)) =>  | 
663  | 
let  | 
|
664  | 
val (haves, have_nots) =  | 
|
665  | 
HOLogic.disjuncts t  | 
|
| 58597 | 666  | 
|> List.partition (exists_subterm (curry (op =) (Var v)))  | 
| 58091 | 667  | 
              |> pairself (fn lits => fold (curry s_disj) lits @{term False})
 | 
668  | 
in  | 
|
669  | 
s_disj (HOLogic.all_const T  | 
|
670  | 
$ Abs (s, T, unskolem (safe_abstract_over (Var v, kill_skolem_arg haves))),  | 
|
671  | 
unskolem have_nots)  | 
|
672  | 
end  | 
|
| 58597 | 673  | 
| NONE =>  | 
674  | 
(case find_var t of  | 
|
675  | 
SOME (v as ((s, _), T)) =>  | 
|
676  | 
HOLogic.all_const T $ Abs (s, T, unskolem (safe_abstract_over (Var v, t)))  | 
|
677  | 
| NONE => t)))  | 
|
| 55195 | 678  | 
in  | 
| 58091 | 679  | 
HOLogic.mk_Trueprop o unskolem o HOLogic.dest_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
 | 
680  | 
end  | 
| 
 
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
 
blanchet 
parents: 
55185 
diff
changeset
 | 
681  | 
|
| 
57788
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
682  | 
fun rename_skolem_args t =  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
683  | 
let  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
684  | 
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
 | 
685  | 
| add_skolem_args t =  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
686  | 
(case strip_comb t of  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
687  | 
(Free (s, _), args as _ :: _) =>  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
688  | 
if String.isPrefix spass_skolem_prefix s then  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
689  | 
insert (op =) (s, fst (take_prefix is_Var args))  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
690  | 
else  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
691  | 
fold add_skolem_args args  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
692  | 
| (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
 | 
693  | 
| _ => I)  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
694  | 
|
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
695  | 
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
 | 
696  | 
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
 | 
697  | 
|
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
698  | 
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
 | 
699  | 
in  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
700  | 
subst_vars ([], subst) t  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
701  | 
end  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
702  | 
|
| 57787 | 703  | 
fun introduce_spass_skolems proof =  | 
704  | 
let  | 
|
| 
57788
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
705  | 
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
 | 
706  | 
| add_skolem _ = I  | 
| 54772 | 707  | 
|
| 57787 | 708  | 
(* union-find would be faster *)  | 
709  | 
fun add_cycle [] = I  | 
|
710  | 
| add_cycle ss =  | 
|
711  | 
fold (fn s => Graph.default_node (s, ())) ss  | 
|
712  | 
#> fold Graph.add_edge (ss ~~ tl ss @ [hd ss])  | 
|
| 54772 | 713  | 
|
| 57787 | 714  | 
val (input_steps, other_steps) = List.partition (null o #5) proof  | 
| 54772 | 715  | 
|
| 58597 | 716  | 
(* The names of the formulas are added to the Skolem constants, to ensure that formulas giving  | 
717  | 
rise to several clauses are skolemized together. *)  | 
|
718  | 
val skoXss = map (fn ((_, ss), _, t, _, _) => Term.fold_aterms add_skolem t ss) input_steps  | 
|
719  | 
val skoXss_input_steps = filter_out (null o fst) (skoXss ~~ input_steps)  | 
|
720  | 
val groups = Graph.strong_conn (fold add_cycle skoXss Graph.empty)  | 
|
| 54772 | 721  | 
|
| 58597 | 722  | 
fun step_name_of_group skoXs = (implode skoXs, [])  | 
| 57787 | 723  | 
fun in_group group = member (op =) group o hd  | 
| 58597 | 724  | 
fun group_of skoX = the (find_first (fn group => in_group group skoX) groups)  | 
| 54772 | 725  | 
|
| 58597 | 726  | 
fun new_steps (skoXss_steps : (string list * (term, 'a) atp_step) list) group =  | 
| 57787 | 727  | 
let  | 
728  | 
val name = step_name_of_group group  | 
|
729  | 
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
 | 
730  | 
val t =  | 
| 58597 | 731  | 
(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
 | 
732  | 
[t] => t  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
733  | 
| ts => ts  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
734  | 
|> 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
 | 
735  | 
|> Library.foldr1 s_conj  | 
| 
 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 
blanchet 
parents: 
57787 
diff
changeset
 | 
736  | 
|> HOLogic.mk_Trueprop)  | 
| 58597 | 737  | 
val skos =  | 
738  | 
fold (union (op =) o filter (String.isPrefix spass_skolem_prefix) o fst) skoXss_steps []  | 
|
739  | 
val deps = map (snd #> #1) skoXss_steps  | 
|
| 57787 | 740  | 
in  | 
741  | 
[(name0, Unknown, unskolemize_term skos t, spass_pre_skolemize_rule, deps),  | 
|
742  | 
(name, Unknown, t, spass_skolemize_rule, [name0])]  | 
|
743  | 
end  | 
|
| 54772 | 744  | 
|
| 57787 | 745  | 
val sko_steps =  | 
| 58597 | 746  | 
maps (fn group => new_steps (filter (in_group group o fst) skoXss_input_steps) group) groups  | 
| 54772 | 747  | 
|
| 57787 | 748  | 
val old_news =  | 
749  | 
map (fn (skos, (name, _, _, _, _)) => (name, [step_name_of_group (group_of skos)]))  | 
|
| 58597 | 750  | 
skoXss_input_steps  | 
| 57787 | 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  | 
|
| 54772 | 755  | 
|
| 57263 | 756  | 
fun factify_atp_proof facts hyp_ts concl_t atp_proof =  | 
| 54505 | 757  | 
let  | 
| 57785 | 758  | 
fun factify_step ((num, ss), role, t, rule, deps) =  | 
| 54505 | 759  | 
let  | 
760  | 
val (ss', role', t') =  | 
|
| 57263 | 761  | 
(case resolve_conjectures ss of  | 
| 54505 | 762  | 
[j] =>  | 
| 
58484
 
b4c0e2b00036
support hypotheses with schematics in Isar proofs
 
blanchet 
parents: 
58477 
diff
changeset
 | 
763  | 
if j = length hyp_ts then ([], Conjecture, concl_t)  | 
| 
 
b4c0e2b00036
support hypotheses with schematics in Isar proofs
 
blanchet 
parents: 
58477 
diff
changeset
 | 
764  | 
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
 | 
765  | 
| _ =>  | 
| 57263 | 766  | 
(case resolve_facts facts ss of  | 
| 57785 | 767  | 
[] => (ss, if role = Lemma then Lemma else Plain, t)  | 
| 
55185
 
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
 
blanchet 
parents: 
54822 
diff
changeset
 | 
768  | 
| facts => (map fst facts, Axiom, t)))  | 
| 54505 | 769  | 
in  | 
770  | 
((num, ss'), role', t', rule, deps)  | 
|
771  | 
end  | 
|
772  | 
||
773  | 
val atp_proof = map factify_step atp_proof  | 
|
774  | 
val names = map #1 atp_proof  | 
|
775  | 
||
776  | 
fun repair_dep (num, ss) = (num, the_default ss (AList.lookup (op =) names num))  | 
|
777  | 
fun repair_deps (name, role, t, rule, deps) = (name, role, t, rule, map repair_dep deps)  | 
|
| 54772 | 778  | 
in  | 
779  | 
map repair_deps atp_proof  | 
|
780  | 
end  | 
|
| 54505 | 781  | 
|
| 31038 | 782  | 
end;  |