| author | wenzelm | 
| Wed, 05 Dec 2012 14:45:44 +0100 | |
| changeset 50366 | b1dd455593a9 | 
| parent 49983 | 33e18e9916a8 | 
| child 50670 | eaa540986291 | 
| 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 | 
| 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: 
45511diff
changeset | 13 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 14 | val metisN : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 15 | val full_typesN : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 16 | val partial_typesN : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 17 | val no_typesN : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 18 | val really_full_type_enc : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 19 | val full_type_enc : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 20 | val partial_type_enc : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 21 | val no_type_enc : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 22 | val full_type_encs : string list | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
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: 
45553diff
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: 
45553diff
changeset | 25 | val metis_call : string -> string -> string | 
| 49983 
33e18e9916a8
use metaquantification when possible in Isar proofs
 blanchet parents: 
49914diff
changeset | 26 | val forall_of : term -> term -> term | 
| 
33e18e9916a8
use metaquantification when possible in Isar proofs
 blanchet parents: 
49914diff
changeset | 27 | val exists_of : term -> term -> term | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
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: 
24387diff
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: 
43062diff
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: 
45511diff
changeset | 45 | val metisN = "metis" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 46 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 47 | val full_typesN = "full_types" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 48 | val partial_typesN = "partial_types" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 49 | val no_typesN = "no_types" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 50 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 51 | val really_full_type_enc = "mono_tags" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 52 | val full_type_enc = "poly_guards_query" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 53 | val partial_type_enc = "poly_args" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 54 | val no_type_enc = "erased" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 55 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 56 | val full_type_encs = [full_type_enc, really_full_type_enc] | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 57 | val partial_type_encs = partial_type_enc :: full_type_encs | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 58 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 59 | val type_enc_aliases = | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 60 | [(full_typesN, full_type_encs), | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 61 | (partial_typesN, partial_type_encs), | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 62 | (no_typesN, [no_type_enc])] | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 63 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 64 | fun unalias_type_enc s = | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 65 | AList.lookup (op =) type_enc_aliases s |> the_default [s] | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 66 | |
| 46365 | 67 | val metis_default_lam_trans = combsN | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 68 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 69 | fun metis_call type_enc lam_trans = | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 70 | let | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 71 | val type_enc = | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 72 | case AList.find (fn (enc, encs) => enc = hd encs) type_enc_aliases | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 73 | type_enc of | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 74 | [alias] => alias | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
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: 
45511diff
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: 
45511diff
changeset | 79 | |
| 39425 | 80 | fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t | 
| 81 | fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda v t | |
| 82 | ||
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 83 | 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 | 84 | 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 | 85 | 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 | 86 | end | 
| 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 87 | |
| 43678 | 88 | exception HO_TERM of (string, string) ho_term list | 
| 48135 | 89 | exception FORMULA of | 
| 90 | (string, string, (string, string) ho_term, string) formula list | |
| 37991 | 91 | exception SAME of unit | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 92 | |
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 93 | (* Type variables are given the basic sort "HOL.type". Some will later be | 
| 37991 | 94 | 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: 
48085diff
changeset | 95 | 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: 
43131diff
changeset | 96 | let val Ts = map (typ_from_atp ctxt) us in | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 97 | case unprefix_and_unascii type_const_prefix a of | 
| 37991 | 98 | SOME b => Type (invert_const b, Ts) | 
| 99 | | NONE => | |
| 100 | if not (null us) then | |
| 43678 | 101 | raise HO_TERM [u] (* only "tconst"s have type arguments *) | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 102 | 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: 
43131diff
changeset | 103 | SOME b => make_tfree ctxt b | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 104 | | 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: 
43296diff
changeset | 105 | (* 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: 
43296diff
changeset | 106 | 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: 
43296diff
changeset | 107 | 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: 
43296diff
changeset | 108 | all cases. *) | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 109 | (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: 
43296diff
changeset | 110 | |> Type_Infer.param 0 | 
| 37991 | 111 | end | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 112 | |
| 38014 | 113 | (* Type class literal applied to a type. Returns triple of polarity, class, | 
| 114 | type. *) | |
| 48132 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 blanchet parents: 
48085diff
changeset | 115 | fun type_constraint_from_term ctxt (u as ATerm ((a, _), us)) = | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 116 | 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: 
42595diff
changeset | 117 | (SOME b, [T]) => (b, T) | 
| 43678 | 118 | | _ => raise HO_TERM [u] | 
| 38014 | 119 | |
| 43907 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 120 | (* Accumulate type constraints in a formula: negative type literals. *) | 
| 38014 | 121 | 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 | 122 | 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 | 123 | | 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 | 124 | | add_type_constraint _ _ = I | 
| 38014 | 125 | |
| 43094 | 126 | fun repair_variable_name f s = | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 127 | let | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 128 | fun subscript_name s n = s ^ nat_subscript n | 
| 38488 | 129 | val s = String.map f s | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 130 | in | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 131 | case space_explode "_" s of | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 132 | [_] => (case take_suffix Char.isDigit (String.explode s) of | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 133 | (cs1 as _ :: _, cs2 as _ :: _) => | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 134 | subscript_name (String.implode cs1) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 135 | (the (Int.fromString (String.implode cs2))) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 136 | | (_, _) => s) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 137 | | [s1, s2] => (case Int.fromString s2 of | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 138 | SOME n => subscript_name s1 n | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 139 | | NONE => s) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 140 | | _ => s | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 141 | end | 
| 43182 
649bada59658
slacker version of "fastype_of", in case a function has dummy type
 blanchet parents: 
43176diff
changeset | 142 | |
| 43907 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 143 | (* 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: 
43863diff
changeset | 144 | (instances of) Skolem pseudoconstants, this information is encoded in the | 
| 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 145 | constant name. *) | 
| 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 146 | fun num_type_args thy s = | 
| 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 147 | 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 | 148 | 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 | 149 | 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 | 150 | 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 | 151 | else | 
| 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 152 | (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 | 153 | |
| 43182 
649bada59658
slacker version of "fastype_of", in case a function has dummy type
 blanchet parents: 
43176diff
changeset | 154 | 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 | 155 | |
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 156 | (* First-order translation. No types are known for variables. "HOLogic.typeT" | 
| 38014 | 157 | 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: 
43131diff
changeset | 158 | fun term_from_atp ctxt textual sym_tab = | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 159 | 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 | 160 | val thy = Proof_Context.theory_of ctxt | 
| 43212 | 161 | (* For Metis, we use 1 rather than 0 because variable references in clauses | 
| 162 | may otherwise conflict with variable constraints in the goal. At least, | |
| 163 | type inference often fails otherwise. See also "axiom_inference" in | |
| 164 | "Metis_Reconstruct". *) | |
| 43094 | 165 | 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 | 166 | fun do_term extra_ts opt_T u = | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 167 | case u of | 
| 48132 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 blanchet parents: 
48085diff
changeset | 168 | ATerm ((s, _), us) => | 
| 46435 | 169 | if String.isPrefix native_type_prefix s then | 
| 42962 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42943diff
changeset | 170 |           @{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 | 171 | else if s = tptp_equal then | 
| 43093 | 172 | let val ts = map (do_term [] NONE) us in | 
| 43094 | 173 | if textual andalso length ts = 2 andalso | 
| 174 | hd ts aconv List.last ts then | |
| 39106 | 175 | (* Vampire is keen on producing these. *) | 
| 176 |               @{const True}
 | |
| 177 | else | |
| 178 |               list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT), ts)
 | |
| 179 | end | |
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 180 | 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: 
44399diff
changeset | 181 | SOME s' => | 
| 42761 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 182 | let | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 183 | val ((s', s''), mangled_us) = | 
| 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 184 | 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: 
42759diff
changeset | 185 | in | 
| 42755 
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
 blanchet parents: 
42751diff
changeset | 186 | 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: 
42587diff
changeset | 187 | 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: 
42587diff
changeset | 188 | [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: 
43131diff
changeset | 189 | do_term extra_ts (SOME (typ_from_atp ctxt typ_u)) term_u | 
| 43678 | 190 | | _ => raise HO_TERM us | 
| 42966 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 blanchet parents: 
42962diff
changeset | 191 | else if s' = predicator_name then | 
| 43093 | 192 |               do_term [] (SOME @{typ bool}) (hd us)
 | 
| 42966 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 blanchet parents: 
42962diff
changeset | 193 | else if s' = app_op_name then | 
| 43131 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 194 | let val extra_t = do_term [] NONE (List.last us) in | 
| 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 195 | do_term (extra_t :: extra_ts) | 
| 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 196 | (case opt_T of | 
| 43182 
649bada59658
slacker version of "fastype_of", in case a function has dummy type
 blanchet parents: 
43176diff
changeset | 197 | SOME T => SOME (slack_fastype_of extra_t --> T) | 
| 43131 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 198 | | NONE => NONE) | 
| 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 199 | (nth us (length us - 2)) | 
| 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 200 | end | 
| 44396 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 blanchet parents: 
44121diff
changeset | 201 | else if s' = type_guard_name then | 
| 42551 
cd99d6d3027a
reconstruct TFF type predicates correctly for ToFoF
 blanchet parents: 
42549diff
changeset | 202 |               @{const True} (* ignore type predicates *)
 | 
| 42549 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 blanchet parents: 
42544diff
changeset | 203 | else | 
| 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 blanchet parents: 
42544diff
changeset | 204 | let | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 205 | 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: 
42751diff
changeset | 206 | 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: 
42751diff
changeset | 207 | 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: 
42544diff
changeset | 208 | val (type_us, term_us) = | 
| 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 blanchet parents: 
42544diff
changeset | 209 | chop num_ty_args us |>> append mangled_us | 
| 43093 | 210 | 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: 
42544diff
changeset | 211 | val T = | 
| 43183 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 212 | (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: 
43182diff
changeset | 213 | num_type_args thy s' = length type_us then | 
| 43191 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 214 | let val Ts = type_us |> map (typ_from_atp ctxt) in | 
| 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 215 | if new_skolem then | 
| 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 216 | SOME (Type_Infer.paramify_vars (tl Ts ---> hd Ts)) | 
| 43200 
ca7b0a48515d
imported patch metis_reconstr_give_type_infer_a_chance
 blanchet parents: 
43199diff
changeset | 217 | else if textual then | 
| 
ca7b0a48515d
imported patch metis_reconstr_give_type_infer_a_chance
 blanchet parents: 
43199diff
changeset | 218 | try (Sign.const_instance thy) (s', Ts) | 
| 43191 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 219 | else | 
| 43200 
ca7b0a48515d
imported patch metis_reconstr_give_type_infer_a_chance
 blanchet parents: 
43199diff
changeset | 220 | NONE | 
| 43191 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 221 | end | 
| 43183 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 222 | else | 
| 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 223 | NONE) | 
| 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 224 | |> (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: 
43182diff
changeset | 225 | | 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: 
43182diff
changeset | 226 | (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: 
43182diff
changeset | 227 | SOME T => T | 
| 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 228 | | NONE => HOLogic.typeT)) | 
| 43191 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 229 | val t = | 
| 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 230 | if new_skolem then | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 231 | Var ((new_skolem_var_name_from_const s'', var_index), T) | 
| 43191 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 232 | else | 
| 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 233 | Const (unproxify_const s', T) | 
| 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 234 | 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: 
42544diff
changeset | 235 | end | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 236 | | NONE => (* a free or schematic variable *) | 
| 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 237 | let | 
| 45042 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44783diff
changeset | 238 | val term_ts = map (do_term [] NONE) us | 
| 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44783diff
changeset | 239 | val ts = term_ts @ extra_ts | 
| 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44783diff
changeset | 240 | val T = | 
| 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44783diff
changeset | 241 | case opt_T of | 
| 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44783diff
changeset | 242 | SOME T => map slack_fastype_of term_ts ---> T | 
| 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44783diff
changeset | 243 | | NONE => map slack_fastype_of ts ---> HOLogic.typeT | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 244 | val t = | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 245 | case unprefix_and_unascii fixed_var_prefix s of | 
| 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 246 | SOME s => Free (s, T) | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 247 | | NONE => | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 248 | 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: 
44399diff
changeset | 249 | SOME s => Var ((s, var_index), T) | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 250 | | NONE => | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 251 | Var ((s |> textual ? repair_variable_name Char.toLower, | 
| 43095 | 252 | var_index), T) | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 253 | in list_comb (t, ts) end | 
| 43093 | 254 | in do_term [] end | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 255 | |
| 48132 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 blanchet parents: 
48085diff
changeset | 256 | fun term_from_atom ctxt textual sym_tab pos (u as ATerm ((s, _), _)) = | 
| 38014 | 257 | 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: 
43131diff
changeset | 258 | add_type_constraint pos (type_constraint_from_term ctxt u) | 
| 38014 | 259 |     #> pair @{const True}
 | 
| 260 | else | |
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 261 |     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: 
36396diff
changeset | 262 | |
| 37991 | 263 | (* Update schematic type variables with detected sort constraints. It's not | 
| 42563 | 264 | totally clear whether this code is necessary. *) | 
| 38014 | 265 | 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 | 266 | let | 
| 37991 | 267 | fun do_type (Type (a, Ts)) = Type (a, map do_type Ts) | 
| 268 | | do_type (TVar (xi, s)) = | |
| 269 | TVar (xi, the_default s (Vartab.lookup tvar_tab xi)) | |
| 270 | | do_type (TFree z) = TFree z | |
| 271 | fun do_term (Const (a, T)) = Const (a, do_type T) | |
| 272 | | do_term (Free (a, T)) = Free (a, do_type T) | |
| 273 | | do_term (Var (xi, T)) = Var (xi, do_type T) | |
| 274 | | do_term (t as Bound _) = t | |
| 275 | | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t) | |
| 276 | | do_term (t1 $ t2) = do_term t1 $ do_term t2 | |
| 277 | in t |> not (Vartab.is_empty tvar_tab) ? do_term end | |
| 278 | ||
| 39425 | 279 | fun quantify_over_var quant_of var_s t = | 
| 280 | let | |
| 281 | val vars = [] |> Term.add_vars t |> filter (fn ((s, _), _) => s = var_s) | |
| 282 | |> map Var | |
| 283 | in fold_rev quant_of vars t end | |
| 37991 | 284 | |
| 38085 
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
 blanchet parents: 
38066diff
changeset | 285 | (* 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: 
38066diff
changeset | 286 | appear in the formula. *) | 
| 43184 | 287 | fun prop_from_atp ctxt textual sym_tab phi = | 
| 38014 | 288 | let | 
| 289 | fun do_formula pos phi = | |
| 37991 | 290 | case phi of | 
| 38014 | 291 | AQuant (_, [], phi) => do_formula pos phi | 
| 42526 | 292 | | AQuant (q, (s, _) :: xs, phi') => | 
| 38014 | 293 | do_formula pos (AQuant (q, xs, phi')) | 
| 42526 | 294 | (* FIXME: TFF *) | 
| 49914 | 295 | #>> quantify_over_var | 
| 296 | (case q of AForall => forall_of | AExists => exists_of) | |
| 297 | (s |> textual ? repair_variable_name Char.toLower) | |
| 38014 | 298 | | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not | 
| 37991 | 299 | | AConn (c, [phi1, phi2]) => | 
| 38014 | 300 | do_formula (pos |> c = AImplies ? not) phi1 | 
| 301 | ##>> do_formula pos phi2 | |
| 302 | #>> (case c of | |
| 303 | AAnd => s_conj | |
| 304 | | AOr => s_disj | |
| 305 | | AImplies => s_imp | |
| 38038 | 306 | | AIff => s_iff | 
| 43163 | 307 | | 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: 
43131diff
changeset | 308 | | AAtom tm => term_from_atom ctxt textual sym_tab pos tm | 
| 37991 | 309 | | _ => raise FORMULA [phi] | 
| 38014 | 310 | in repair_tvar_sorts (do_formula true phi Vartab.empty) end | 
| 37991 | 311 | |
| 31038 | 312 | end; |