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