src/HOL/Tools/ATP/atp_proof_reconstruct.ML
author blanchet
Tue, 19 Nov 2013 19:36:24 +0100
changeset 54505 d023838eb252
parent 54501 77c9460e01b0
child 54506 8b5caa190054
permissions -rw-r--r--
more refactoring to accommodate SMT proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45887
diff changeset
     1
(*  Title:      HOL/Tools/ATP/atp_proof_reconstruct.ML
38027
505657ddb047 standardize "Author" tags
blanchet
parents: 38019
diff changeset
     2
    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
505657ddb047 standardize "Author" tags
blanchet
parents: 38019
diff changeset
     3
    Author:     Claire Quigley, Cambridge University Computer Laboratory
36392
c00c57850eb7 cosmetics: rename internal functions
blanchet
parents: 36369
diff changeset
     4
    Author:     Jasmin Blanchette, TU Muenchen
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
     5
49914
23e36a4d28f1 refactor code
blanchet
parents: 49881
diff changeset
     6
Basic proof reconstruction from ATP proofs.
33310
44f9665c2091 proper header;
wenzelm
parents: 33243
diff changeset
     7
*)
44f9665c2091 proper header;
wenzelm
parents: 33243
diff changeset
     8
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45887
diff changeset
     9
signature ATP_PROOF_RECONSTRUCT =
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    10
sig
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
    11
  type ('a, 'b) atp_term = ('a, 'b) ATP_Problem.atp_term
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
    12
  type ('a, 'b, 'c, 'd) atp_formula = ('a, 'b, 'c, 'd) ATP_Problem.atp_formula
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    13
  type stature = ATP_Problem_Generate.stature
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
    14
  type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
    15
  type 'a atp_proof = 'a ATP_Proof.atp_proof
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    16
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    17
  val metisN : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    18
  val full_typesN : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    19
  val partial_typesN : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    20
  val no_typesN : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    21
  val really_full_type_enc : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    22
  val full_type_enc : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    23
  val partial_type_enc : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    24
  val no_type_enc : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    25
  val full_type_encs : string list
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    26
  val partial_type_encs : string list
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    27
  val default_metis_lam_trans : string
45554
09ad83de849c don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents: 45553
diff changeset
    28
  val metis_call : string -> string -> string
49983
33e18e9916a8 use metaquantification when possible in Isar proofs
blanchet
parents: 49914
diff changeset
    29
  val forall_of : term -> term -> term
33e18e9916a8 use metaquantification when possible in Isar proofs
blanchet
parents: 49914
diff changeset
    30
  val exists_of : term -> term -> term
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    31
  val unalias_type_enc : string -> string list
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
    32
  val term_of_atp :
49914
23e36a4d28f1 refactor code
blanchet
parents: 49881
diff changeset
    33
    Proof.context -> bool -> int Symtab.table -> typ option ->
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
    34
    (string, string) atp_term -> term
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
    35
  val prop_of_atp :
49914
23e36a4d28f1 refactor code
blanchet
parents: 49881
diff changeset
    36
    Proof.context -> bool -> int Symtab.table ->
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
    37
    (string, string, (string, string) atp_term, string) atp_formula -> term
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
    38
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    39
  val used_facts_in_atp_proof :
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    40
    Proof.context -> (string * stature) list vector -> string atp_proof ->
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    41
    (string * stature) list
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    42
  val used_facts_in_unsound_atp_proof :
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    43
    Proof.context -> (string * stature) list vector -> 'a atp_proof ->
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    44
    string list option
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    45
  val lam_trans_of_atp_proof : string atp_proof -> string -> string
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    46
  val is_typed_helper_used_in_atp_proof : string atp_proof -> bool
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
    47
  val termify_atp_proof : Proof.context -> string Symtab.table -> (string * term) list ->
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
    48
    int Symtab.table -> string atp_proof -> (term, string) atp_step list
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
    49
  val factify_atp_proof : (string * 'a) list vector -> term list -> term ->
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
    50
    (term, string) atp_step list -> (term, string) atp_step list
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    51
end;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    52
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45887
diff changeset
    53
structure ATP_Proof_Reconstruct : ATP_PROOF_RECONSTRUCT =
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    54
struct
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    55
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43062
diff changeset
    56
open ATP_Util
38028
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
    57
open ATP_Problem
39452
70a57e40f795 factored out TSTP/SPASS/Vampire proof parsing;
blanchet
parents: 39425
diff changeset
    58
open ATP_Proof
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45887
diff changeset
    59
open ATP_Problem_Generate
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    60
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    61
val metisN = "metis"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    62
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    63
val full_typesN = "full_types"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    64
val partial_typesN = "partial_types"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    65
val no_typesN = "no_types"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    66
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    67
val really_full_type_enc = "mono_tags"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    68
val full_type_enc = "poly_guards_query"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    69
val partial_type_enc = "poly_args"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    70
val no_type_enc = "erased"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    71
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    72
val full_type_encs = [full_type_enc, really_full_type_enc]
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    73
val partial_type_encs = partial_type_enc :: full_type_encs
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    74
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    75
val type_enc_aliases =
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    76
  [(full_typesN, full_type_encs),
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    77
   (partial_typesN, partial_type_encs),
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    78
   (no_typesN, [no_type_enc])]
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    79
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    80
fun unalias_type_enc s =
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    81
  AList.lookup (op =) type_enc_aliases s |> the_default [s]
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    82
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    83
val default_metis_lam_trans = combsN
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    84
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    85
fun metis_call type_enc lam_trans =
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    86
  let
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    87
    val type_enc =
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    88
      case AList.find (fn (enc, encs) => enc = hd encs) type_enc_aliases
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    89
                      type_enc of
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    90
        [alias] => alias
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    91
      | _ => type_enc
45522
3b951bbd2bee compile
blanchet
parents: 45520
diff changeset
    92
    val opts = [] |> type_enc <> partial_typesN ? cons type_enc
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    93
                  |> lam_trans <> default_metis_lam_trans ? cons lam_trans
45520
2b1dde0b1c30 thread in additional options to minimizer
blanchet
parents: 45519
diff changeset
    94
  in metisN ^ (if null opts then "" else " (" ^ commas opts ^ ")") end
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    95
50670
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 49983
diff changeset
    96
fun term_name' (Var ((s, _), _)) = perhaps (try Name.dest_skolem) s
52034
11b48e7a4e7e correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents: 52031
diff changeset
    97
  | term_name' _ = ""
50670
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 49983
diff changeset
    98
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 49983
diff changeset
    99
fun lambda' v = Term.lambda_name (term_name' v, v)
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 49983
diff changeset
   100
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 49983
diff changeset
   101
fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda' v t
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 49983
diff changeset
   102
fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda' v t
39425
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   103
43135
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   104
fun make_tfree ctxt w =
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   105
  let val ww = "'" ^ w in
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   106
    TFree (ww, the_default HOLogic.typeS (Variable.def_sort ctxt (ww, ~1)))
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   107
  end
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   108
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
   109
exception ATP_TERM of (string, string) atp_term list
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
   110
exception ATP_FORMULA of
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
   111
    (string, string, (string, string) atp_term, string) atp_formula list
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   112
exception SAME of unit
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   113
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   114
(* Type variables are given the basic sort "HOL.type". Some will later be
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   115
   constrained by information from type literals, or by type inference. *)
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   116
fun typ_of_atp ctxt (u as ATerm ((a, _), us)) =
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   117
  let val Ts = map (typ_of_atp ctxt) us in
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45379
diff changeset
   118
    case unprefix_and_unascii type_const_prefix a of
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   119
      SOME b => Type (invert_const b, Ts)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   120
    | NONE =>
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   121
      if not (null us) then
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
   122
        raise ATP_TERM [u]  (* only "tconst"s have type arguments *)
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45379
diff changeset
   123
      else case unprefix_and_unascii tfree_prefix a of
43135
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   124
        SOME b => make_tfree ctxt b
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   125
      | NONE =>
43302
566f970006e5 improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents: 43296
diff changeset
   126
        (* Could be an Isabelle variable or a variable from the ATP, say "X1"
566f970006e5 improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents: 43296
diff changeset
   127
           or "_5018". Sometimes variables from the ATP are indistinguishable
566f970006e5 improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents: 43296
diff changeset
   128
           from Isabelle variables, which forces us to use a type parameter in
566f970006e5 improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents: 43296
diff changeset
   129
           all cases. *)
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45379
diff changeset
   130
        (a |> perhaps (unprefix_and_unascii tvar_prefix), HOLogic.typeS)
43302
566f970006e5 improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents: 43296
diff changeset
   131
        |> Type_Infer.param 0
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   132
  end
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   133
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   134
(* Type class literal applied to a type. Returns triple of polarity, class,
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   135
   type. *)
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   136
fun type_constraint_of_term ctxt (u as ATerm ((a, _), us)) =
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   137
  case (unprefix_and_unascii class_prefix a, map (typ_of_atp ctxt) us) of
42606
0c76cf483899 show sorts not just types in Isar proofs + tuning
blanchet
parents: 42595
diff changeset
   138
    (SOME b, [T]) => (b, T)
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
   139
  | _ => raise ATP_TERM [u]
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   140
43907
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   141
(* Accumulate type constraints in a formula: negative type literals. *)
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   142
fun add_var (key, z)  = Vartab.map_default (key, []) (cons z)
42606
0c76cf483899 show sorts not just types in Isar proofs + tuning
blanchet
parents: 42595
diff changeset
   143
fun add_type_constraint false (cl, TFree (a ,_)) = add_var ((a, ~1), cl)
0c76cf483899 show sorts not just types in Isar proofs + tuning
blanchet
parents: 42595
diff changeset
   144
  | add_type_constraint false (cl, TVar (ix, _)) = add_var (ix, cl)
0c76cf483899 show sorts not just types in Isar proofs + tuning
blanchet
parents: 42595
diff changeset
   145
  | add_type_constraint _ _ = I
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   146
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50670
diff changeset
   147
fun repair_var_name s =
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   148
  let
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   149
    fun subscript_name s n = s ^ nat_subscript n
50704
cd1fcda1ea88 rename variable in binder, not just in body
blanchet
parents: 50703
diff changeset
   150
    val s = s |> String.map Char.toLower
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   151
  in
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   152
    case space_explode "_" s of
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   153
      [_] => (case take_suffix Char.isDigit (String.explode s) of
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   154
                (cs1 as _ :: _, cs2 as _ :: _) =>
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   155
                subscript_name (String.implode cs1)
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   156
                               (the (Int.fromString (String.implode cs2)))
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   157
              | (_, _) => s)
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   158
    | [s1, s2] => (case Int.fromString s2 of
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   159
                     SOME n => subscript_name s1 n
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   160
                   | NONE => s)
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   161
    | _ => s
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   162
  end
43182
649bada59658 slacker version of "fastype_of", in case a function has dummy type
blanchet
parents: 43176
diff changeset
   163
43907
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   164
(* The number of type arguments of a constant, zero if it's monomorphic. For
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   165
   (instances of) Skolem pseudoconstants, this information is encoded in the
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   166
   constant name. *)
52125
ac7830871177 improved handling of free variables' types in Isar proofs
blanchet
parents: 52034
diff changeset
   167
fun robust_const_num_type_args thy s =
43907
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   168
  if String.isPrefix skolem_const_prefix s then
46711
f745bcc4a1e5 more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
wenzelm
parents: 46435
diff changeset
   169
    s |> Long_Name.explode |> List.last |> Int.fromString |> the
45554
09ad83de849c don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents: 45553
diff changeset
   170
  else if String.isPrefix lam_lifted_prefix s then
09ad83de849c don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents: 45553
diff changeset
   171
    if String.isPrefix lam_lifted_poly_prefix s then 2 else 0
43907
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   172
  else
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   173
    (s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   174
43182
649bada59658 slacker version of "fastype_of", in case a function has dummy type
blanchet
parents: 43176
diff changeset
   175
fun slack_fastype_of t = fastype_of t handle TERM _ => HOLogic.typeT
649bada59658 slacker version of "fastype_of", in case a function has dummy type
blanchet
parents: 43176
diff changeset
   176
51192
4dc6bb65c3c3 slacker comparison for Skolems, to avoid trivial equations
blanchet
parents: 50704
diff changeset
   177
(* Cope with "tt(X) = X" atoms, where "X" is existentially quantified. *)
4dc6bb65c3c3 slacker comparison for Skolems, to avoid trivial equations
blanchet
parents: 50704
diff changeset
   178
fun loose_aconv (Free (s, _), Free (s', _)) = s = s'
4dc6bb65c3c3 slacker comparison for Skolems, to avoid trivial equations
blanchet
parents: 50704
diff changeset
   179
  | loose_aconv (t, t') = t aconv t'
4dc6bb65c3c3 slacker comparison for Skolems, to avoid trivial equations
blanchet
parents: 50704
diff changeset
   180
50703
76a2e506c125 swap Vampire's Skolem arguments to bring them in line with what E and metis's new skolemizer do (helps Isar proof reconstruction in some cases)
blanchet
parents: 50693
diff changeset
   181
val vampire_skolem_prefix = "sK"
76a2e506c125 swap Vampire's Skolem arguments to bring them in line with what E and metis's new skolemizer do (helps Isar proof reconstruction in some cases)
blanchet
parents: 50693
diff changeset
   182
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   183
(* First-order translation. No types are known for variables. "HOLogic.typeT"
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   184
   should allow them to be inferred. *)
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   185
fun term_of_atp ctxt textual sym_tab =
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   186
  let
43135
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   187
    val thy = Proof_Context.theory_of ctxt
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43204
diff changeset
   188
    (* For Metis, we use 1 rather than 0 because variable references in clauses
050a03afe024 Metis code cleanup
blanchet
parents: 43204
diff changeset
   189
       may otherwise conflict with variable constraints in the goal. At least,
050a03afe024 Metis code cleanup
blanchet
parents: 43204
diff changeset
   190
       type inference often fails otherwise. See also "axiom_inference" in
050a03afe024 Metis code cleanup
blanchet
parents: 43204
diff changeset
   191
       "Metis_Reconstruct". *)
43094
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
   192
    val var_index = if textual then 0 else 1
43131
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   193
    fun do_term extra_ts opt_T u =
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   194
      case u of
48132
9aa0fad4e864 added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents: 48085
diff changeset
   195
        ATerm ((s, _), us) =>
51878
f11039b31bae handle dummy atp terms
smolkas
parents: 51192
diff changeset
   196
        if s = ""
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
   197
          then error "Isar proof reconstruction failed because the ATP proof \
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
   198
                     \contains unparsable material."
51878
f11039b31bae handle dummy atp terms
smolkas
parents: 51192
diff changeset
   199
        else if String.isPrefix native_type_prefix s then
42962
3b50fdeb6cfc started adding support for THF output (but no lambdas)
blanchet
parents: 42943
diff changeset
   200
          @{const True} (* ignore TPTP type information *)
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   201
        else if s = tptp_equal then
43093
blanchet
parents: 43085
diff changeset
   202
          let val ts = map (do_term [] NONE) us in
43094
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
   203
            if textual andalso length ts = 2 andalso
51192
4dc6bb65c3c3 slacker comparison for Skolems, to avoid trivial equations
blanchet
parents: 50704
diff changeset
   204
               loose_aconv (hd ts, List.last ts) then
39106
5ab6a3707499 fix trivial "x = x" fact detection
blanchet
parents: 38988
diff changeset
   205
              @{const True}
5ab6a3707499 fix trivial "x = x" fact detection
blanchet
parents: 38988
diff changeset
   206
            else
5ab6a3707499 fix trivial "x = x" fact detection
blanchet
parents: 38988
diff changeset
   207
              list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT), ts)
5ab6a3707499 fix trivial "x = x" fact detection
blanchet
parents: 38988
diff changeset
   208
          end
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45379
diff changeset
   209
        else case unprefix_and_unascii const_prefix s of
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   210
          SOME s' =>
42761
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   211
          let
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   212
            val ((s', s''), mangled_us) =
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   213
              s' |> unmangled_const |>> `invert_const
42761
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   214
          in
42755
4603154a3018 robustly detect how many type args were passed to the ATP, even if some of them were omitted
blanchet
parents: 42751
diff changeset
   215
            if s' = type_tag_name then
42589
9f7c48463645 restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents: 42587
diff changeset
   216
              case mangled_us @ us of
9f7c48463645 restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents: 42587
diff changeset
   217
                [typ_u, term_u] =>
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   218
                do_term extra_ts (SOME (typ_of_atp ctxt typ_u)) term_u
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
   219
              | _ => raise ATP_TERM us
42966
4e2d6c1e5392 more work on parsing LEO-II proofs without lambdas
blanchet
parents: 42962
diff changeset
   220
            else if s' = predicator_name then
43093
blanchet
parents: 43085
diff changeset
   221
              do_term [] (SOME @{typ bool}) (hd us)
42966
4e2d6c1e5392 more work on parsing LEO-II proofs without lambdas
blanchet
parents: 42962
diff changeset
   222
            else if s' = app_op_name then
43131
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   223
              let val extra_t = do_term [] NONE (List.last us) in
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   224
                do_term (extra_t :: extra_ts)
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   225
                        (case opt_T of
43182
649bada59658 slacker version of "fastype_of", in case a function has dummy type
blanchet
parents: 43176
diff changeset
   226
                           SOME T => SOME (slack_fastype_of extra_t --> T)
43131
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   227
                         | NONE => NONE)
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   228
                        (nth us (length us - 2))
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   229
              end
44396
66b9b3fcd608 started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents: 44121
diff changeset
   230
            else if s' = type_guard_name then
42551
cd99d6d3027a reconstruct TFF type predicates correctly for ToFoF
blanchet
parents: 42549
diff changeset
   231
              @{const True} (* ignore type predicates *)
42549
b9754f26c7bc handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents: 42544
diff changeset
   232
            else
b9754f26c7bc handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents: 42544
diff changeset
   233
              let
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   234
                val new_skolem = String.isPrefix new_skolem_const_prefix s''
42755
4603154a3018 robustly detect how many type args were passed to the ATP, even if some of them were omitted
blanchet
parents: 42751
diff changeset
   235
                val num_ty_args =
4603154a3018 robustly detect how many type args were passed to the ATP, even if some of them were omitted
blanchet
parents: 42751
diff changeset
   236
                  length us - the_default 0 (Symtab.lookup sym_tab s)
42549
b9754f26c7bc handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents: 42544
diff changeset
   237
                val (type_us, term_us) =
b9754f26c7bc handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents: 42544
diff changeset
   238
                  chop num_ty_args us |>> append mangled_us
43093
blanchet
parents: 43085
diff changeset
   239
                val term_ts = map (do_term [] NONE) term_us
42549
b9754f26c7bc handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents: 42544
diff changeset
   240
                val T =
43183
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   241
                  (if not (null type_us) andalso
52125
ac7830871177 improved handling of free variables' types in Isar proofs
blanchet
parents: 52034
diff changeset
   242
                      robust_const_num_type_args thy s' = length type_us then
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   243
                     let val Ts = type_us |> map (typ_of_atp ctxt) in
43191
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   244
                       if new_skolem then
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   245
                         SOME (Type_Infer.paramify_vars (tl Ts ---> hd Ts))
43200
ca7b0a48515d imported patch metis_reconstr_give_type_infer_a_chance
blanchet
parents: 43199
diff changeset
   246
                       else if textual then
ca7b0a48515d imported patch metis_reconstr_give_type_infer_a_chance
blanchet
parents: 43199
diff changeset
   247
                         try (Sign.const_instance thy) (s', Ts)
43191
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   248
                       else
43200
ca7b0a48515d imported patch metis_reconstr_give_type_infer_a_chance
blanchet
parents: 43199
diff changeset
   249
                         NONE
43191
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   250
                     end
43183
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   251
                   else
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   252
                     NONE)
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   253
                  |> (fn SOME T => T
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   254
                       | NONE => map slack_fastype_of term_ts --->
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   255
                                 (case opt_T of
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   256
                                    SOME T => T
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   257
                                  | NONE => HOLogic.typeT))
43191
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   258
                val t =
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   259
                  if new_skolem then
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   260
                    Var ((new_skolem_var_name_of_const s'', var_index), T)
43191
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   261
                  else
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   262
                    Const (unproxify_const s', T)
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   263
              in list_comb (t, term_ts @ extra_ts) end
42549
b9754f26c7bc handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents: 42544
diff changeset
   264
          end
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   265
        | NONE => (* a free or schematic variable *)
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   266
          let
50693
2fac44deb8b5 tuned comment
blanchet
parents: 50676
diff changeset
   267
            (* This assumes that distinct names are mapped to distinct names by
2fac44deb8b5 tuned comment
blanchet
parents: 50676
diff changeset
   268
               "Variable.variant_frees". This does not hold in general but
2fac44deb8b5 tuned comment
blanchet
parents: 50676
diff changeset
   269
               should hold for ATP-generated Skolem function names, since these
2fac44deb8b5 tuned comment
blanchet
parents: 50676
diff changeset
   270
               end with a digit and "variant_frees" appends letters. *)
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50670
diff changeset
   271
            fun fresh_up s =
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50670
diff changeset
   272
              [(s, ())] |> Variable.variant_frees ctxt [] |> hd |> fst
50703
76a2e506c125 swap Vampire's Skolem arguments to bring them in line with what E and metis's new skolemizer do (helps Isar proof reconstruction in some cases)
blanchet
parents: 50693
diff changeset
   273
            val term_ts =
76a2e506c125 swap Vampire's Skolem arguments to bring them in line with what E and metis's new skolemizer do (helps Isar proof reconstruction in some cases)
blanchet
parents: 50693
diff changeset
   274
              map (do_term [] NONE) us
76a2e506c125 swap Vampire's Skolem arguments to bring them in line with what E and metis's new skolemizer do (helps Isar proof reconstruction in some cases)
blanchet
parents: 50693
diff changeset
   275
              (* Vampire (2.6) passes arguments to Skolem functions in reverse
76a2e506c125 swap Vampire's Skolem arguments to bring them in line with what E and metis's new skolemizer do (helps Isar proof reconstruction in some cases)
blanchet
parents: 50693
diff changeset
   276
                 order *)
76a2e506c125 swap Vampire's Skolem arguments to bring them in line with what E and metis's new skolemizer do (helps Isar proof reconstruction in some cases)
blanchet
parents: 50693
diff changeset
   277
              |> String.isPrefix vampire_skolem_prefix s ? rev
45042
89341b897412 better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents: 44783
diff changeset
   278
            val ts = term_ts @ extra_ts
89341b897412 better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents: 44783
diff changeset
   279
            val T =
89341b897412 better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents: 44783
diff changeset
   280
              case opt_T of
89341b897412 better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents: 44783
diff changeset
   281
                SOME T => map slack_fastype_of term_ts ---> T
89341b897412 better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents: 44783
diff changeset
   282
              | NONE => map slack_fastype_of ts ---> HOLogic.typeT
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   283
            val t =
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45379
diff changeset
   284
              case unprefix_and_unascii fixed_var_prefix s of
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45379
diff changeset
   285
                SOME s => Free (s, T)
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   286
              | NONE =>
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45379
diff changeset
   287
                case unprefix_and_unascii schematic_var_prefix s of
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   288
                  SOME s => Var ((s, var_index), T)
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   289
                | NONE =>
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50670
diff changeset
   290
                  if textual andalso not (is_tptp_variable s) then
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50670
diff changeset
   291
                    Free (s |> textual ? (repair_var_name #> fresh_up), T)
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50670
diff changeset
   292
                  else
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50670
diff changeset
   293
                    Var ((s |> textual ? repair_var_name, var_index), T)
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   294
          in list_comb (t, ts) end
43093
blanchet
parents: 43085
diff changeset
   295
  in do_term [] end
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   296
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   297
fun term_of_atom ctxt textual sym_tab pos (u as ATerm ((s, _), _)) =
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   298
  if String.isPrefix class_prefix s then
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   299
    add_type_constraint pos (type_constraint_of_term ctxt u)
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   300
    #> pair @{const True}
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   301
  else
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   302
    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: 36396
diff changeset
   303
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   304
(* Update schematic type variables with detected sort constraints. It's not
42563
e70ffe3846d0 improve helper type instantiation code
blanchet
parents: 42562
diff changeset
   305
   totally clear whether this code is necessary. *)
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   306
fun repair_tvar_sorts (t, tvar_tab) =
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   307
  let
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   308
    fun do_type (Type (a, Ts)) = Type (a, map do_type Ts)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   309
      | do_type (TVar (xi, s)) =
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   310
        TVar (xi, the_default s (Vartab.lookup tvar_tab xi))
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   311
      | do_type (TFree z) = TFree z
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   312
    fun do_term (Const (a, T)) = Const (a, do_type T)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   313
      | do_term (Free (a, T)) = Free (a, do_type T)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   314
      | do_term (Var (xi, T)) = Var (xi, do_type T)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   315
      | do_term (t as Bound _) = t
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   316
      | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   317
      | do_term (t1 $ t2) = do_term t1 $ do_term t2
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   318
  in t |> not (Vartab.is_empty tvar_tab) ? do_term end
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   319
39425
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   320
fun quantify_over_var quant_of var_s t =
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   321
  let
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   322
    val vars = [] |> Term.add_vars t |> filter (fn ((s, _), _) => s = var_s)
52758
7ffcd6f2890d avoid duplicating Var when the types do not quite fit -- since this step occurs before type inference
blanchet
parents: 52125
diff changeset
   323
    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: 52125
diff changeset
   324
    fun norm_var_types (Var (x, T)) =
7ffcd6f2890d avoid duplicating Var when the types do not quite fit -- since this step occurs before type inference
blanchet
parents: 52125
diff changeset
   325
        Var (x, case AList.lookup (op =) normTs x of
7ffcd6f2890d avoid duplicating Var when the types do not quite fit -- since this step occurs before type inference
blanchet
parents: 52125
diff changeset
   326
                  NONE => T
7ffcd6f2890d avoid duplicating Var when the types do not quite fit -- since this step occurs before type inference
blanchet
parents: 52125
diff changeset
   327
                | SOME T' => T')
7ffcd6f2890d avoid duplicating Var when the types do not quite fit -- since this step occurs before type inference
blanchet
parents: 52125
diff changeset
   328
      | 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: 52125
diff changeset
   329
  in t |> map_aterms norm_var_types |> fold_rev quant_of (map Var normTs) end
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   330
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38066
diff changeset
   331
(* Interpret an ATP formula as a HOL term, extracting sort constraints as they
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38066
diff changeset
   332
   appear in the formula. *)
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   333
fun prop_of_atp ctxt textual sym_tab phi =
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   334
  let
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   335
    fun do_formula pos phi =
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   336
      case phi of
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   337
        AQuant (_, [], phi) => do_formula pos phi
42526
46d485f8d144 added room for types in ATP quantifiers
blanchet
parents: 42451
diff changeset
   338
      | AQuant (q, (s, _) :: xs, phi') =>
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   339
        do_formula pos (AQuant (q, xs, phi'))
42526
46d485f8d144 added room for types in ATP quantifiers
blanchet
parents: 42451
diff changeset
   340
        (* FIXME: TFF *)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49881
diff changeset
   341
        #>> quantify_over_var
23e36a4d28f1 refactor code
blanchet
parents: 49881
diff changeset
   342
              (case q of AForall => forall_of | AExists => exists_of)
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50670
diff changeset
   343
              (s |> textual ? repair_var_name)
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   344
      | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   345
      | AConn (c, [phi1, phi2]) =>
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   346
        do_formula (pos |> c = AImplies ? not) phi1
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   347
        ##>> do_formula pos phi2
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   348
        #>> (case c of
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   349
               AAnd => s_conj
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   350
             | AOr => s_disj
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   351
             | AImplies => s_imp
38038
584ab1a3a523 more robust proof reconstruction
blanchet
parents: 38036
diff changeset
   352
             | AIff => s_iff
43163
31babd4b1552 killed odd connectives
blanchet
parents: 43136
diff changeset
   353
             | ANot => raise Fail "impossible connective")
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   354
      | AAtom tm => term_of_atom ctxt textual sym_tab pos tm
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
   355
      | _ => raise ATP_FORMULA [phi]
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   356
  in repair_tvar_sorts (do_formula true phi Vartab.empty) end
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   357
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   358
fun find_first_in_list_vector vec key =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   359
  Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   360
                 | (_, value) => value) NONE vec
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   361
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   362
val unprefix_fact_number = space_implode "_" o tl o space_explode "_"
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   363
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   364
fun resolve_one_named_fact fact_names s =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   365
  case try (unprefix fact_prefix) s of
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   366
    SOME s' =>
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   367
    let val s' = s' |> unprefix_fact_number |> unascii_of in
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   368
      s' |> find_first_in_list_vector fact_names |> Option.map (pair s')
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   369
    end
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   370
  | NONE => NONE
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   371
fun resolve_fact fact_names = map_filter (resolve_one_named_fact fact_names)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   372
fun is_fact fact_names = not o null o resolve_fact fact_names
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   373
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   374
fun resolve_one_named_conjecture s =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   375
  case try (unprefix conjecture_prefix) s of
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   376
    SOME s' => Int.fromString s'
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   377
  | NONE => NONE
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   378
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   379
val resolve_conjecture = map_filter resolve_one_named_conjecture
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   380
val is_conjecture = not o null o resolve_conjecture
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   381
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   382
fun is_axiom_used_in_proof pred =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   383
  exists (fn ((_, ss), _, _, _, []) => exists pred ss | _ => false)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   384
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   385
fun add_non_rec_defs fact_names accum =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   386
  Vector.foldl (fn (facts, facts') =>
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   387
      union (op =) (filter (fn (_, (_, status)) => status = Non_Rec_Def) facts)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   388
            facts')
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   389
    accum fact_names
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   390
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   391
val isa_ext = Thm.get_name_hint @{thm ext}
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   392
val isa_short_ext = Long_Name.base_name isa_ext
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   393
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   394
fun ext_name ctxt =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   395
  if Thm.eq_thm_prop (@{thm ext},
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   396
       singleton (Attrib.eval_thms ctxt) (Facts.named isa_short_ext, [])) then
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   397
    isa_short_ext
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   398
  else
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   399
    isa_ext
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   400
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   401
val leo2_extcnf_equal_neg_rule = "extcnf_equal_neg"
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   402
val leo2_unfold_def_rule = "unfold_def"
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   403
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   404
fun add_fact ctxt fact_names ((_, ss), _, _, rule, deps) =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   405
  (if rule = leo2_extcnf_equal_neg_rule then
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   406
     insert (op =) (ext_name ctxt, (Global, General))
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   407
   else if rule = leo2_unfold_def_rule then
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   408
     (* LEO 1.3.3 does not record definitions properly, leading to missing
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   409
        dependencies in the TSTP proof. Remove the next line once this is
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   410
        fixed. *)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   411
     add_non_rec_defs fact_names
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   412
   else if rule = agsyhol_coreN orelse rule = satallax_coreN then
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   413
     (fn [] =>
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   414
         (* agsyHOL and Satallax don't include definitions in their
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   415
            unsatisfiable cores, so we assume the worst and include them all
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   416
            here. *)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   417
         [(ext_name ctxt, (Global, General))] |> add_non_rec_defs fact_names
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   418
       | facts => facts)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   419
   else
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   420
     I)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   421
  #> (if null deps then union (op =) (resolve_fact fact_names ss) else I)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   422
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   423
fun used_facts_in_atp_proof ctxt fact_names atp_proof =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   424
  if null atp_proof then Vector.foldl (uncurry (union (op =))) [] fact_names
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   425
  else fold (add_fact ctxt fact_names) atp_proof []
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   426
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   427
fun used_facts_in_unsound_atp_proof _ _ [] = NONE
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   428
  | used_facts_in_unsound_atp_proof ctxt fact_names atp_proof =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   429
    let val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof in
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   430
      if forall (fn (_, (sc, _)) => sc = Global) used_facts andalso
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   431
         not (is_axiom_used_in_proof (is_conjecture o single) atp_proof) then
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   432
        SOME (map fst used_facts)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   433
      else
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   434
        NONE
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   435
    end
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   436
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   437
val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   438
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   439
(* overapproximation (good enough) *)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   440
fun is_lam_lifted s =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   441
  String.isPrefix fact_prefix s andalso
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   442
  String.isSubstring ascii_of_lam_fact_prefix s
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   443
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   444
val is_combinator_def = String.isPrefix (helper_prefix ^ combinator_prefix)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   445
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   446
fun lam_trans_of_atp_proof atp_proof default =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   447
  case (is_axiom_used_in_proof is_combinator_def atp_proof,
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   448
        is_axiom_used_in_proof is_lam_lifted atp_proof) of
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   449
    (false, false) => default
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   450
  | (false, true) => liftingN
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   451
(*  | (true, true) => combs_and_liftingN -- not supported by "metis" *)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   452
  | (true, _) => combsN
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   453
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   454
val is_typed_helper_name =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   455
  String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   456
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   457
fun is_typed_helper_used_in_atp_proof atp_proof =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   458
  is_axiom_used_in_proof is_typed_helper_name atp_proof
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   459
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   460
fun repair_name "$true" = "c_True"
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   461
  | repair_name "$false" = "c_False"
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   462
  | repair_name "$$e" = tptp_equal (* seen in Vampire proofs *)
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   463
  | repair_name s =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   464
    if is_tptp_equal s orelse
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   465
       (* seen in Vampire proofs *)
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   466
       (String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   467
      tptp_equal
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   468
    else
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   469
      s
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   470
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   471
fun infer_formula_types ctxt =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   472
  Type.constraint HOLogic.boolT
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   473
  #> Syntax.check_term
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   474
         (Proof_Context.set_mode Proof_Context.mode_schematic ctxt)
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   475
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   476
val combinator_table =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   477
  [(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def [abs_def]}),
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   478
   (@{const_name Meson.COMBK}, @{thm Meson.COMBK_def [abs_def]}),
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   479
   (@{const_name Meson.COMBB}, @{thm Meson.COMBB_def [abs_def]}),
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   480
   (@{const_name Meson.COMBC}, @{thm Meson.COMBC_def [abs_def]}),
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   481
   (@{const_name Meson.COMBS}, @{thm Meson.COMBS_def [abs_def]})]
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   482
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   483
fun uncombine_term thy =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   484
  let
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   485
    fun aux (t1 $ t2) = betapply (pairself aux (t1, t2))
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   486
      | aux (Abs (s, T, t')) = Abs (s, T, aux t')
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   487
      | aux (t as Const (x as (s, _))) =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   488
        (case AList.lookup (op =) combinator_table s of
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   489
           SOME thm => thm |> prop_of |> specialize_type thy x
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   490
                           |> Logic.dest_equals |> snd
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   491
         | NONE => t)
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   492
      | aux t = t
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   493
  in aux end
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   494
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   495
fun unlift_term lifted =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   496
  map_aterms (fn t as Const (s, _) =>
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   497
                 if String.isPrefix lam_lifted_prefix s then
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   498
                   case AList.lookup (op =) lifted s of
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   499
                     SOME t =>
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   500
                     (* FIXME: do something about the types *)
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   501
                     unlift_term lifted t
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   502
                   | NONE => t
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   503
                 else
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   504
                   t
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   505
               | t => t)
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   506
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   507
fun decode_line ctxt lifted sym_tab (name, role, u, rule, deps) =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   508
  let
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   509
    val thy = Proof_Context.theory_of ctxt
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   510
    val t =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   511
      u |> prop_of_atp ctxt true sym_tab
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   512
        |> uncombine_term thy
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   513
        |> unlift_term lifted
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   514
        |> infer_formula_types ctxt
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   515
  in (name, role, t, rule, deps) end
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   516
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   517
val waldmeister_conjecture_num = "1.0.0.0"
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   518
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   519
fun repair_waldmeister_endgame arg =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   520
  let
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   521
    fun do_tail (name, _, t, rule, deps) =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   522
      (name, Negated_Conjecture, s_not t, rule, deps)
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   523
    fun do_body [] = []
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   524
      | do_body ((line as ((num, _), _, _, _, _)) :: lines) =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   525
        if num = waldmeister_conjecture_num then map do_tail (line :: lines)
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   526
        else line :: do_body lines
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   527
  in do_body arg end
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   528
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   529
fun termify_atp_proof ctxt pool lifted sym_tab =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   530
  clean_up_atp_proof_dependencies
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   531
  #> nasty_atp_proof pool
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   532
  #> map_term_names_in_atp_proof repair_name
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   533
  #> map (decode_line ctxt lifted sym_tab)
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   534
  #> repair_waldmeister_endgame
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   535
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   536
fun factify_atp_proof fact_names hyp_ts concl_t atp_proof =
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   537
  let
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   538
    fun factify_step ((num, ss), role, t, rule, deps) =
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   539
      let
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   540
        val (ss', role', t') =
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   541
          (case resolve_conjecture ss of
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   542
            [j] =>
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   543
            if j = length hyp_ts then ([], Conjecture, concl_t) else ([], Hypothesis, nth hyp_ts j)
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   544
           | _ =>
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   545
             (case resolve_fact fact_names ss of
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   546
               [] => (ss, Plain, t)
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   547
             | facts => (map fst facts, Axiom, t)))
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   548
      in
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   549
        ((num, ss'), role', t', rule, deps)
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   550
      end
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   551
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   552
    val atp_proof = map factify_step atp_proof
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   553
    val names = map #1 atp_proof
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   554
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   555
    fun repair_dep (num, ss) = (num, the_default ss (AList.lookup (op =) names num))
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   556
    fun repair_deps (name, role, t, rule, deps) = (name, role, t, rule, map repair_dep deps)
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   557
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   558
  in map repair_deps atp_proof end
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   559
31038
immler@in.tum.de
parents: 31037
diff changeset
   560
end;