| author | blanchet | 
| Wed, 20 Feb 2013 08:44:24 +0100 | |
| changeset 51192 | 4dc6bb65c3c3 | 
| parent 50704 | cd1fcda1ea88 | 
| child 51878 | f11039b31bae | 
| 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  | 
| 43678 | 11  | 
  type ('a, 'b) ho_term = ('a, 'b) ATP_Problem.ho_term
 | 
| 48135 | 12  | 
  type ('a, 'b, 'c, 'd) formula = ('a, 'b, 'c, 'd) ATP_Problem.formula
 | 
| 
45519
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
13  | 
|
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
14  | 
val metisN : string  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
15  | 
val full_typesN : string  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
16  | 
val partial_typesN : string  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
17  | 
val no_typesN : string  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
18  | 
val really_full_type_enc : string  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
19  | 
val full_type_enc : string  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
20  | 
val partial_type_enc : string  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
21  | 
val no_type_enc : string  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
22  | 
val full_type_encs : string list  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
23  | 
val partial_type_encs : string list  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45553 
diff
changeset
 | 
24  | 
val metis_default_lam_trans : string  | 
| 
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45553 
diff
changeset
 | 
25  | 
val metis_call : string -> string -> string  | 
| 
49983
 
33e18e9916a8
use metaquantification when possible in Isar proofs
 
blanchet 
parents: 
49914 
diff
changeset
 | 
26  | 
val forall_of : term -> term -> term  | 
| 
 
33e18e9916a8
use metaquantification when possible in Isar proofs
 
blanchet 
parents: 
49914 
diff
changeset
 | 
27  | 
val exists_of : term -> term -> term  | 
| 
45519
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
28  | 
val unalias_type_enc : string -> string list  | 
| 43094 | 29  | 
val term_from_atp :  | 
| 49914 | 30  | 
Proof.context -> bool -> int Symtab.table -> typ option ->  | 
31  | 
(string, string) ho_term -> term  | 
|
| 43127 | 32  | 
val prop_from_atp :  | 
| 49914 | 33  | 
Proof.context -> bool -> int Symtab.table ->  | 
34  | 
(string, string, (string, string) ho_term, string) formula -> term  | 
|
| 
24425
 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 
paulson 
parents: 
24387 
diff
changeset
 | 
35  | 
end;  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
36  | 
|
| 46320 | 37  | 
structure ATP_Proof_Reconstruct : ATP_PROOF_RECONSTRUCT =  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
38  | 
struct  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
39  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43062 
diff
changeset
 | 
40  | 
open ATP_Util  | 
| 38028 | 41  | 
open ATP_Problem  | 
| 39452 | 42  | 
open ATP_Proof  | 
| 46320 | 43  | 
open ATP_Problem_Generate  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
44  | 
|
| 
45519
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
45  | 
val metisN = "metis"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
46  | 
|
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
47  | 
val full_typesN = "full_types"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
48  | 
val partial_typesN = "partial_types"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
49  | 
val no_typesN = "no_types"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
50  | 
|
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
51  | 
val really_full_type_enc = "mono_tags"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
52  | 
val full_type_enc = "poly_guards_query"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
53  | 
val partial_type_enc = "poly_args"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
54  | 
val no_type_enc = "erased"  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
55  | 
|
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
56  | 
val full_type_encs = [full_type_enc, really_full_type_enc]  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
57  | 
val partial_type_encs = partial_type_enc :: full_type_encs  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
58  | 
|
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
59  | 
val type_enc_aliases =  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
60  | 
[(full_typesN, full_type_encs),  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
61  | 
(partial_typesN, partial_type_encs),  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
62  | 
(no_typesN, [no_type_enc])]  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
63  | 
|
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
64  | 
fun unalias_type_enc s =  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
65  | 
AList.lookup (op =) type_enc_aliases s |> the_default [s]  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
66  | 
|
| 46365 | 67  | 
val metis_default_lam_trans = combsN  | 
| 
45519
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
68  | 
|
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
69  | 
fun metis_call type_enc lam_trans =  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
70  | 
let  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
71  | 
val type_enc =  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
72  | 
case AList.find (fn (enc, encs) => enc = hd encs) type_enc_aliases  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
73  | 
type_enc of  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
74  | 
[alias] => alias  | 
| 
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
75  | 
| _ => type_enc  | 
| 45522 | 76  | 
val opts = [] |> type_enc <> partial_typesN ? cons type_enc  | 
| 
45519
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
77  | 
|> lam_trans <> metis_default_lam_trans ? cons lam_trans  | 
| 45520 | 78  | 
  in metisN ^ (if null opts then "" else " (" ^ commas opts ^ ")") end
 | 
| 
45519
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45511 
diff
changeset
 | 
79  | 
|
| 
50670
 
eaa540986291
properly take the existential closure of skolems
 
blanchet 
parents: 
49983 
diff
changeset
 | 
80  | 
fun term_name' (Var ((s, _), _)) = perhaps (try Name.dest_skolem) s  | 
| 
 
eaa540986291
properly take the existential closure of skolems
 
blanchet 
parents: 
49983 
diff
changeset
 | 
81  | 
| term_name' t = ""  | 
| 
 
eaa540986291
properly take the existential closure of skolems
 
blanchet 
parents: 
49983 
diff
changeset
 | 
82  | 
|
| 
 
eaa540986291
properly take the existential closure of skolems
 
blanchet 
parents: 
49983 
diff
changeset
 | 
83  | 
fun lambda' v = Term.lambda_name (term_name' v, v)  | 
| 
 
eaa540986291
properly take the existential closure of skolems
 
blanchet 
parents: 
49983 
diff
changeset
 | 
84  | 
|
| 
 
eaa540986291
properly take the existential closure of skolems
 
blanchet 
parents: 
49983 
diff
changeset
 | 
85  | 
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
 | 
86  | 
fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda' v t  | 
| 39425 | 87  | 
|
| 
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
 | 
88  | 
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
 | 
89  | 
let val ww = "'" ^ w in  | 
| 
 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 
blanchet 
parents: 
43131 
diff
changeset
 | 
90  | 
TFree (ww, the_default HOLogic.typeS (Variable.def_sort ctxt (ww, ~1)))  | 
| 
 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 
blanchet 
parents: 
43131 
diff
changeset
 | 
91  | 
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
 | 
92  | 
|
| 43678 | 93  | 
exception HO_TERM of (string, string) ho_term list  | 
| 48135 | 94  | 
exception FORMULA of  | 
95  | 
(string, string, (string, string) ho_term, string) formula list  | 
|
| 37991 | 96  | 
exception SAME of unit  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
97  | 
|
| 
36909
 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 
blanchet 
parents: 
36607 
diff
changeset
 | 
98  | 
(* Type variables are given the basic sort "HOL.type". Some will later be  | 
| 37991 | 99  | 
constrained by information from type literals, or by type inference. *)  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48085 
diff
changeset
 | 
100  | 
fun typ_from_atp ctxt (u as ATerm ((a, _), us)) =  | 
| 
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
 | 
101  | 
let val Ts = map (typ_from_atp ctxt) us in  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45379 
diff
changeset
 | 
102  | 
case unprefix_and_unascii type_const_prefix a of  | 
| 37991 | 103  | 
SOME b => Type (invert_const b, Ts)  | 
104  | 
| NONE =>  | 
|
105  | 
if not (null us) then  | 
|
| 43678 | 106  | 
raise HO_TERM [u] (* only "tconst"s have type arguments *)  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45379 
diff
changeset
 | 
107  | 
else case unprefix_and_unascii tfree_prefix a of  | 
| 
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
 | 
108  | 
SOME b => make_tfree ctxt b  | 
| 
36486
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
109  | 
| NONE =>  | 
| 
43302
 
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
 
blanchet 
parents: 
43296 
diff
changeset
 | 
110  | 
(* Could be an Isabelle variable or a variable from the ATP, say "X1"  | 
| 
 
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
 
blanchet 
parents: 
43296 
diff
changeset
 | 
111  | 
or "_5018". Sometimes variables from the ATP are indistinguishable  | 
| 
 
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
 
blanchet 
parents: 
43296 
diff
changeset
 | 
112  | 
from Isabelle variables, which forces us to use a type parameter in  | 
| 
 
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
 
blanchet 
parents: 
43296 
diff
changeset
 | 
113  | 
all cases. *)  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45379 
diff
changeset
 | 
114  | 
(a |> perhaps (unprefix_and_unascii tvar_prefix), HOLogic.typeS)  | 
| 
43302
 
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
 
blanchet 
parents: 
43296 
diff
changeset
 | 
115  | 
|> Type_Infer.param 0  | 
| 37991 | 116  | 
end  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
117  | 
|
| 38014 | 118  | 
(* Type class literal applied to a type. Returns triple of polarity, class,  | 
119  | 
type. *)  | 
|
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48085 
diff
changeset
 | 
120  | 
fun type_constraint_from_term ctxt (u as ATerm ((a, _), us)) =  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45379 
diff
changeset
 | 
121  | 
case (unprefix_and_unascii class_prefix a, map (typ_from_atp ctxt) us) of  | 
| 
42606
 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 
blanchet 
parents: 
42595 
diff
changeset
 | 
122  | 
(SOME b, [T]) => (b, T)  | 
| 43678 | 123  | 
| _ => raise HO_TERM [u]  | 
| 38014 | 124  | 
|
| 
43907
 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 
blanchet 
parents: 
43863 
diff
changeset
 | 
125  | 
(* Accumulate type constraints in a formula: negative type literals. *)  | 
| 38014 | 126  | 
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
 | 
127  | 
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
 | 
128  | 
| 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
 | 
129  | 
| add_type_constraint _ _ = I  | 
| 38014 | 130  | 
|
| 
50676
 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 
blanchet 
parents: 
50670 
diff
changeset
 | 
131  | 
fun repair_var_name s =  | 
| 
36486
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
132  | 
let  | 
| 
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
133  | 
fun subscript_name s n = s ^ nat_subscript n  | 
| 50704 | 134  | 
val s = s |> String.map Char.toLower  | 
| 
36486
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
135  | 
in  | 
| 
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
136  | 
case space_explode "_" s of  | 
| 
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
137  | 
[_] => (case take_suffix Char.isDigit (String.explode s) of  | 
| 
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
138  | 
(cs1 as _ :: _, cs2 as _ :: _) =>  | 
| 
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
139  | 
subscript_name (String.implode cs1)  | 
| 
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
140  | 
(the (Int.fromString (String.implode cs2)))  | 
| 
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
141  | 
| (_, _) => s)  | 
| 
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
142  | 
| [s1, s2] => (case Int.fromString s2 of  | 
| 
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
143  | 
SOME n => subscript_name s1 n  | 
| 
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
144  | 
| NONE => s)  | 
| 
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
145  | 
| _ => s  | 
| 
 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 
blanchet 
parents: 
36485 
diff
changeset
 | 
146  | 
end  | 
| 
43182
 
649bada59658
slacker version of "fastype_of", in case a function has dummy type
 
blanchet 
parents: 
43176 
diff
changeset
 | 
147  | 
|
| 
43907
 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 
blanchet 
parents: 
43863 
diff
changeset
 | 
148  | 
(* The number of type arguments of a constant, zero if it's monomorphic. For  | 
| 
 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 
blanchet 
parents: 
43863 
diff
changeset
 | 
149  | 
(instances of) Skolem pseudoconstants, this information is encoded in the  | 
| 
 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 
blanchet 
parents: 
43863 
diff
changeset
 | 
150  | 
constant name. *)  | 
| 
 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 
blanchet 
parents: 
43863 
diff
changeset
 | 
151  | 
fun num_type_args thy s =  | 
| 
 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 
blanchet 
parents: 
43863 
diff
changeset
 | 
152  | 
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
 | 
153  | 
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
 | 
154  | 
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
 | 
155  | 
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
 | 
156  | 
else  | 
| 
 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 
blanchet 
parents: 
43863 
diff
changeset
 | 
157  | 
(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
 | 
158  | 
|
| 
43182
 
649bada59658
slacker version of "fastype_of", in case a function has dummy type
 
blanchet 
parents: 
43176 
diff
changeset
 | 
159  | 
fun slack_fastype_of t = fastype_of t handle TERM _ => HOLogic.typeT  | 
| 
 
649bada59658
slacker version of "fastype_of", in case a function has dummy type
 
blanchet 
parents: 
43176 
diff
changeset
 | 
160  | 
|
| 
51192
 
4dc6bb65c3c3
slacker comparison for Skolems, to avoid trivial equations
 
blanchet 
parents: 
50704 
diff
changeset
 | 
161  | 
(* Cope with "tt(X) = X" atoms, where "X" is existentially quantified. *)  | 
| 
 
4dc6bb65c3c3
slacker comparison for Skolems, to avoid trivial equations
 
blanchet 
parents: 
50704 
diff
changeset
 | 
162  | 
fun loose_aconv (Free (s, _), Free (s', _)) = s = s'  | 
| 
 
4dc6bb65c3c3
slacker comparison for Skolems, to avoid trivial equations
 
blanchet 
parents: 
50704 
diff
changeset
 | 
163  | 
| loose_aconv (t, t') = t aconv t'  | 
| 
 
4dc6bb65c3c3
slacker comparison for Skolems, to avoid trivial equations
 
blanchet 
parents: 
50704 
diff
changeset
 | 
164  | 
|
| 
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
 | 
165  | 
val vampire_skolem_prefix = "sK"  | 
| 
 
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
 | 
166  | 
|
| 
36909
 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 
blanchet 
parents: 
36607 
diff
changeset
 | 
167  | 
(* First-order translation. No types are known for variables. "HOLogic.typeT"  | 
| 38014 | 168  | 
should allow them to be inferred. *)  | 
| 
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
 | 
169  | 
fun term_from_atp ctxt textual sym_tab =  | 
| 
36909
 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 
blanchet 
parents: 
36607 
diff
changeset
 | 
170  | 
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
 | 
171  | 
val thy = Proof_Context.theory_of ctxt  | 
| 43212 | 172  | 
(* For Metis, we use 1 rather than 0 because variable references in clauses  | 
173  | 
may otherwise conflict with variable constraints in the goal. At least,  | 
|
174  | 
type inference often fails otherwise. See also "axiom_inference" in  | 
|
175  | 
"Metis_Reconstruct". *)  | 
|
| 43094 | 176  | 
val var_index = if textual then 0 else 1  | 
| 
43131
 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 
blanchet 
parents: 
43130 
diff
changeset
 | 
177  | 
fun do_term extra_ts opt_T u =  | 
| 
36909
 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 
blanchet 
parents: 
36607 
diff
changeset
 | 
178  | 
case u of  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48085 
diff
changeset
 | 
179  | 
ATerm ((s, _), us) =>  | 
| 46435 | 180  | 
if String.isPrefix native_type_prefix s then  | 
| 
42962
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42943 
diff
changeset
 | 
181  | 
          @{const True} (* ignore TPTP type information *)
 | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44399 
diff
changeset
 | 
182  | 
else if s = tptp_equal then  | 
| 43093 | 183  | 
let val ts = map (do_term [] NONE) us in  | 
| 43094 | 184  | 
if textual andalso length ts = 2 andalso  | 
| 
51192
 
4dc6bb65c3c3
slacker comparison for Skolems, to avoid trivial equations
 
blanchet 
parents: 
50704 
diff
changeset
 | 
185  | 
loose_aconv (hd ts, List.last ts) then  | 
| 39106 | 186  | 
              @{const True}
 | 
187  | 
else  | 
|
188  | 
              list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT), ts)
 | 
|
189  | 
end  | 
|
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45379 
diff
changeset
 | 
190  | 
else case unprefix_and_unascii const_prefix s of  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44399 
diff
changeset
 | 
191  | 
SOME s' =>  | 
| 
42761
 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 
blanchet 
parents: 
42759 
diff
changeset
 | 
192  | 
let  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44399 
diff
changeset
 | 
193  | 
val ((s', s''), mangled_us) =  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44399 
diff
changeset
 | 
194  | 
s' |> unmangled_const |>> `invert_const  | 
| 
42761
 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 
blanchet 
parents: 
42759 
diff
changeset
 | 
195  | 
in  | 
| 
42755
 
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
 
blanchet 
parents: 
42751 
diff
changeset
 | 
196  | 
if s' = type_tag_name then  | 
| 
42589
 
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
 
blanchet 
parents: 
42587 
diff
changeset
 | 
197  | 
case mangled_us @ us of  | 
| 
 
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
 
blanchet 
parents: 
42587 
diff
changeset
 | 
198  | 
[typ_u, term_u] =>  | 
| 
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
 | 
199  | 
do_term extra_ts (SOME (typ_from_atp ctxt typ_u)) term_u  | 
| 43678 | 200  | 
| _ => raise HO_TERM us  | 
| 
42966
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42962 
diff
changeset
 | 
201  | 
else if s' = predicator_name then  | 
| 43093 | 202  | 
              do_term [] (SOME @{typ bool}) (hd us)
 | 
| 
42966
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42962 
diff
changeset
 | 
203  | 
else if s' = app_op_name then  | 
| 
43131
 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 
blanchet 
parents: 
43130 
diff
changeset
 | 
204  | 
let val extra_t = do_term [] NONE (List.last us) in  | 
| 
 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 
blanchet 
parents: 
43130 
diff
changeset
 | 
205  | 
do_term (extra_t :: extra_ts)  | 
| 
 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 
blanchet 
parents: 
43130 
diff
changeset
 | 
206  | 
(case opt_T of  | 
| 
43182
 
649bada59658
slacker version of "fastype_of", in case a function has dummy type
 
blanchet 
parents: 
43176 
diff
changeset
 | 
207  | 
SOME T => SOME (slack_fastype_of extra_t --> T)  | 
| 
43131
 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 
blanchet 
parents: 
43130 
diff
changeset
 | 
208  | 
| NONE => NONE)  | 
| 
 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 
blanchet 
parents: 
43130 
diff
changeset
 | 
209  | 
(nth us (length us - 2))  | 
| 
 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 
blanchet 
parents: 
43130 
diff
changeset
 | 
210  | 
end  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44121 
diff
changeset
 | 
211  | 
else if s' = type_guard_name then  | 
| 
42551
 
cd99d6d3027a
reconstruct TFF type predicates correctly for ToFoF
 
blanchet 
parents: 
42549 
diff
changeset
 | 
212  | 
              @{const True} (* ignore type predicates *)
 | 
| 
42549
 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 
blanchet 
parents: 
42544 
diff
changeset
 | 
213  | 
else  | 
| 
 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 
blanchet 
parents: 
42544 
diff
changeset
 | 
214  | 
let  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44399 
diff
changeset
 | 
215  | 
val new_skolem = String.isPrefix new_skolem_const_prefix s''  | 
| 
42755
 
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
 
blanchet 
parents: 
42751 
diff
changeset
 | 
216  | 
val num_ty_args =  | 
| 
 
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
 
blanchet 
parents: 
42751 
diff
changeset
 | 
217  | 
length us - the_default 0 (Symtab.lookup sym_tab s)  | 
| 
42549
 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 
blanchet 
parents: 
42544 
diff
changeset
 | 
218  | 
val (type_us, term_us) =  | 
| 
 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 
blanchet 
parents: 
42544 
diff
changeset
 | 
219  | 
chop num_ty_args us |>> append mangled_us  | 
| 43093 | 220  | 
val term_ts = map (do_term [] NONE) term_us  | 
| 
42549
 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 
blanchet 
parents: 
42544 
diff
changeset
 | 
221  | 
val T =  | 
| 
43183
 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 
blanchet 
parents: 
43182 
diff
changeset
 | 
222  | 
(if not (null type_us) andalso  | 
| 
 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 
blanchet 
parents: 
43182 
diff
changeset
 | 
223  | 
num_type_args thy s' = length type_us then  | 
| 
43191
 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 
blanchet 
parents: 
43184 
diff
changeset
 | 
224  | 
let val Ts = type_us |> map (typ_from_atp ctxt) in  | 
| 
 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 
blanchet 
parents: 
43184 
diff
changeset
 | 
225  | 
if new_skolem then  | 
| 
 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 
blanchet 
parents: 
43184 
diff
changeset
 | 
226  | 
SOME (Type_Infer.paramify_vars (tl Ts ---> hd Ts))  | 
| 
43200
 
ca7b0a48515d
imported patch metis_reconstr_give_type_infer_a_chance
 
blanchet 
parents: 
43199 
diff
changeset
 | 
227  | 
else if textual then  | 
| 
 
ca7b0a48515d
imported patch metis_reconstr_give_type_infer_a_chance
 
blanchet 
parents: 
43199 
diff
changeset
 | 
228  | 
try (Sign.const_instance thy) (s', Ts)  | 
| 
43191
 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 
blanchet 
parents: 
43184 
diff
changeset
 | 
229  | 
else  | 
| 
43200
 
ca7b0a48515d
imported patch metis_reconstr_give_type_infer_a_chance
 
blanchet 
parents: 
43199 
diff
changeset
 | 
230  | 
NONE  | 
| 
43191
 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 
blanchet 
parents: 
43184 
diff
changeset
 | 
231  | 
end  | 
| 
43183
 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 
blanchet 
parents: 
43182 
diff
changeset
 | 
232  | 
else  | 
| 
 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 
blanchet 
parents: 
43182 
diff
changeset
 | 
233  | 
NONE)  | 
| 
 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 
blanchet 
parents: 
43182 
diff
changeset
 | 
234  | 
|> (fn SOME T => T  | 
| 
 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 
blanchet 
parents: 
43182 
diff
changeset
 | 
235  | 
| NONE => map slack_fastype_of term_ts --->  | 
| 
 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 
blanchet 
parents: 
43182 
diff
changeset
 | 
236  | 
(case opt_T of  | 
| 
 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 
blanchet 
parents: 
43182 
diff
changeset
 | 
237  | 
SOME T => T  | 
| 
 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 
blanchet 
parents: 
43182 
diff
changeset
 | 
238  | 
| NONE => HOLogic.typeT))  | 
| 
43191
 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 
blanchet 
parents: 
43184 
diff
changeset
 | 
239  | 
val t =  | 
| 
 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 
blanchet 
parents: 
43184 
diff
changeset
 | 
240  | 
if new_skolem then  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44399 
diff
changeset
 | 
241  | 
Var ((new_skolem_var_name_from_const s'', var_index), T)  | 
| 
43191
 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 
blanchet 
parents: 
43184 
diff
changeset
 | 
242  | 
else  | 
| 
 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 
blanchet 
parents: 
43184 
diff
changeset
 | 
243  | 
Const (unproxify_const s', T)  | 
| 
 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 
blanchet 
parents: 
43184 
diff
changeset
 | 
244  | 
in list_comb (t, term_ts @ extra_ts) end  | 
| 
42549
 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 
blanchet 
parents: 
42544 
diff
changeset
 | 
245  | 
end  | 
| 
36909
 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 
blanchet 
parents: 
36607 
diff
changeset
 | 
246  | 
| NONE => (* a free or schematic variable *)  | 
| 
 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 
blanchet 
parents: 
36607 
diff
changeset
 | 
247  | 
let  | 
| 50693 | 248  | 
(* This assumes that distinct names are mapped to distinct names by  | 
249  | 
"Variable.variant_frees". This does not hold in general but  | 
|
250  | 
should hold for ATP-generated Skolem function names, since these  | 
|
251  | 
end with a digit and "variant_frees" appends letters. *)  | 
|
| 
50676
 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 
blanchet 
parents: 
50670 
diff
changeset
 | 
252  | 
fun fresh_up s =  | 
| 
 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 
blanchet 
parents: 
50670 
diff
changeset
 | 
253  | 
[(s, ())] |> Variable.variant_frees ctxt [] |> hd |> fst  | 
| 
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
 | 
254  | 
val term_ts =  | 
| 
 
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
 | 
255  | 
map (do_term [] NONE) us  | 
| 
 
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
 | 
256  | 
(* Vampire (2.6) passes arguments to Skolem functions in reverse  | 
| 
 
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
 | 
257  | 
order *)  | 
| 
 
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
 | 
258  | 
|> String.isPrefix vampire_skolem_prefix s ? rev  | 
| 
45042
 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 
blanchet 
parents: 
44783 
diff
changeset
 | 
259  | 
val ts = term_ts @ extra_ts  | 
| 
 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 
blanchet 
parents: 
44783 
diff
changeset
 | 
260  | 
val T =  | 
| 
 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 
blanchet 
parents: 
44783 
diff
changeset
 | 
261  | 
case opt_T of  | 
| 
 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 
blanchet 
parents: 
44783 
diff
changeset
 | 
262  | 
SOME T => map slack_fastype_of term_ts ---> T  | 
| 
 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 
blanchet 
parents: 
44783 
diff
changeset
 | 
263  | 
| NONE => map slack_fastype_of ts ---> HOLogic.typeT  | 
| 
36909
 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 
blanchet 
parents: 
36607 
diff
changeset
 | 
264  | 
val t =  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45379 
diff
changeset
 | 
265  | 
case unprefix_and_unascii fixed_var_prefix s of  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45379 
diff
changeset
 | 
266  | 
SOME s => Free (s, T)  | 
| 
36909
 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 
blanchet 
parents: 
36607 
diff
changeset
 | 
267  | 
| NONE =>  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45379 
diff
changeset
 | 
268  | 
case unprefix_and_unascii schematic_var_prefix s of  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44399 
diff
changeset
 | 
269  | 
SOME s => Var ((s, var_index), T)  | 
| 
36909
 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 
blanchet 
parents: 
36607 
diff
changeset
 | 
270  | 
| NONE =>  | 
| 
50676
 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 
blanchet 
parents: 
50670 
diff
changeset
 | 
271  | 
if textual andalso not (is_tptp_variable s) then  | 
| 
 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 
blanchet 
parents: 
50670 
diff
changeset
 | 
272  | 
Free (s |> textual ? (repair_var_name #> fresh_up), T)  | 
| 
 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 
blanchet 
parents: 
50670 
diff
changeset
 | 
273  | 
else  | 
| 
 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 
blanchet 
parents: 
50670 
diff
changeset
 | 
274  | 
Var ((s |> textual ? repair_var_name, var_index), T)  | 
| 
36909
 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 
blanchet 
parents: 
36607 
diff
changeset
 | 
275  | 
in list_comb (t, ts) end  | 
| 43093 | 276  | 
in do_term [] end  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
277  | 
|
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48085 
diff
changeset
 | 
278  | 
fun term_from_atom ctxt textual sym_tab pos (u as ATerm ((s, _), _)) =  | 
| 38014 | 279  | 
if String.isPrefix class_prefix s then  | 
| 
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
 | 
280  | 
add_type_constraint pos (type_constraint_from_term ctxt u)  | 
| 38014 | 281  | 
    #> pair @{const True}
 | 
282  | 
else  | 
|
| 
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
 | 
283  | 
    pair (term_from_atp ctxt 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
 | 
284  | 
|
| 37991 | 285  | 
(* Update schematic type variables with detected sort constraints. It's not  | 
| 42563 | 286  | 
totally clear whether this code is necessary. *)  | 
| 38014 | 287  | 
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
 | 
288  | 
let  | 
| 37991 | 289  | 
fun do_type (Type (a, Ts)) = Type (a, map do_type Ts)  | 
290  | 
| do_type (TVar (xi, s)) =  | 
|
291  | 
TVar (xi, the_default s (Vartab.lookup tvar_tab xi))  | 
|
292  | 
| do_type (TFree z) = TFree z  | 
|
293  | 
fun do_term (Const (a, T)) = Const (a, do_type T)  | 
|
294  | 
| do_term (Free (a, T)) = Free (a, do_type T)  | 
|
295  | 
| do_term (Var (xi, T)) = Var (xi, do_type T)  | 
|
296  | 
| do_term (t as Bound _) = t  | 
|
297  | 
| do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t)  | 
|
298  | 
| do_term (t1 $ t2) = do_term t1 $ do_term t2  | 
|
299  | 
in t |> not (Vartab.is_empty tvar_tab) ? do_term end  | 
|
300  | 
||
| 39425 | 301  | 
fun quantify_over_var quant_of var_s t =  | 
302  | 
let  | 
|
303  | 
val vars = [] |> Term.add_vars t |> filter (fn ((s, _), _) => s = var_s)  | 
|
304  | 
|> map Var  | 
|
305  | 
in fold_rev quant_of vars t end  | 
|
| 37991 | 306  | 
|
| 
38085
 
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
 
blanchet 
parents: 
38066 
diff
changeset
 | 
307  | 
(* Interpret an ATP formula as a HOL term, extracting sort constraints as they  | 
| 
 
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
 
blanchet 
parents: 
38066 
diff
changeset
 | 
308  | 
appear in the formula. *)  | 
| 43184 | 309  | 
fun prop_from_atp ctxt textual sym_tab phi =  | 
| 38014 | 310  | 
let  | 
311  | 
fun do_formula pos phi =  | 
|
| 37991 | 312  | 
case phi of  | 
| 38014 | 313  | 
AQuant (_, [], phi) => do_formula pos phi  | 
| 42526 | 314  | 
| AQuant (q, (s, _) :: xs, phi') =>  | 
| 38014 | 315  | 
do_formula pos (AQuant (q, xs, phi'))  | 
| 42526 | 316  | 
(* FIXME: TFF *)  | 
| 49914 | 317  | 
#>> quantify_over_var  | 
318  | 
(case q of AForall => forall_of | AExists => exists_of)  | 
|
| 
50676
 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 
blanchet 
parents: 
50670 
diff
changeset
 | 
319  | 
(s |> textual ? repair_var_name)  | 
| 38014 | 320  | 
| AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not  | 
| 37991 | 321  | 
| AConn (c, [phi1, phi2]) =>  | 
| 38014 | 322  | 
do_formula (pos |> c = AImplies ? not) phi1  | 
323  | 
##>> do_formula pos phi2  | 
|
324  | 
#>> (case c of  | 
|
325  | 
AAnd => s_conj  | 
|
326  | 
| AOr => s_disj  | 
|
327  | 
| AImplies => s_imp  | 
|
| 38038 | 328  | 
| AIff => s_iff  | 
| 43163 | 329  | 
| ANot => raise Fail "impossible connective")  | 
| 
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  | 
| AAtom tm => term_from_atom ctxt textual sym_tab pos tm  | 
| 37991 | 331  | 
| _ => raise FORMULA [phi]  | 
| 38014 | 332  | 
in repair_tvar_sorts (do_formula true phi Vartab.empty) end  | 
| 37991 | 333  | 
|
| 31038 | 334  | 
end;  |