src/HOL/Tools/ATP/atp_proof_reconstruct.ML
author fleury
Wed, 30 Jul 2014 14:03:11 +0200
changeset 57703 fefe3ea75289
parent 57699 a6cf197c1f1e
child 57707 0242e9578828
permissions -rw-r--r--
removing the '= True' generated by Leo-II.
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
54811
df56a01f5684 parse SPASS-Pirate types
blanchet
parents: 54799
diff changeset
    11
  type 'a atp_type = 'a ATP_Problem.atp_type
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
    12
  type ('a, 'b) atp_term = ('a, 'b) ATP_Problem.atp_term
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
    13
  type ('a, 'b, 'c, 'd) atp_formula = ('a, 'b, 'c, 'd) ATP_Problem.atp_formula
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    14
  type stature = ATP_Problem_Generate.stature
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
    15
  type atp_step_name = ATP_Proof.atp_step_name
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
    16
  type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
    17
  type 'a atp_proof = 'a ATP_Proof.atp_proof
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    18
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    19
  val metisN : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    20
  val full_typesN : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    21
  val partial_typesN : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    22
  val no_typesN : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    23
  val really_full_type_enc : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    24
  val full_type_enc : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    25
  val partial_type_enc : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    26
  val no_type_enc : string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    27
  val full_type_encs : string list
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    28
  val partial_type_encs : string list
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    29
  val default_metis_lam_trans : string
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
    30
49983
33e18e9916a8 use metaquantification when possible in Isar proofs
blanchet
parents: 49914
diff changeset
    31
  val forall_of : term -> term -> term
33e18e9916a8 use metaquantification when possible in Isar proofs
blanchet
parents: 49914
diff changeset
    32
  val exists_of : term -> term -> term
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55257
diff changeset
    33
  val type_enc_aliases : (string * string list) list
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    34
  val unalias_type_enc : string -> string list
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
    35
  val term_of_atp : Proof.context -> ATP_Problem.atp_format -> ATP_Problem_Generate.type_enc ->
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
    36
    bool -> int Symtab.table -> typ option -> (string, string atp_type) atp_term -> term
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
    37
  val prop_of_atp : Proof.context -> ATP_Problem.atp_format -> ATP_Problem_Generate.type_enc ->
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
    38
    bool -> int Symtab.table ->
54811
df56a01f5684 parse SPASS-Pirate types
blanchet
parents: 54799
diff changeset
    39
    (string, string, (string, string atp_type) atp_term, string) atp_formula -> term
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
    40
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
    41
  val used_facts_in_atp_proof : Proof.context -> (string * stature) list -> string atp_proof ->
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
    42
    (string * stature) list
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
    43
  val used_facts_in_unsound_atp_proof : Proof.context -> (string * stature) list -> 'a atp_proof ->
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
    44
    string list option
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55195
diff changeset
    45
  val atp_proof_prefers_lifting : string atp_proof -> bool
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    46
  val is_typed_helper_used_in_atp_proof : string atp_proof -> bool
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
    47
  val replace_dependencies_in_line : atp_step_name * atp_step_name list -> ('a, 'b) atp_step ->
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
    48
    ('a, 'b) atp_step
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
    49
  val termify_atp_proof : Proof.context -> string -> ATP_Problem.atp_format ->
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
    50
    ATP_Problem_Generate.type_enc -> string Symtab.table -> (string * term) list ->
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
    51
    int Symtab.table -> string atp_proof -> (term, string) atp_step list
57267
8b87114357bd integrated new Waldmeister code with 'sledgehammer' command
blanchet
parents: 57263
diff changeset
    52
  val repair_waldmeister_endgame : (term, 'a) atp_step list -> (term, 'a) atp_step list
57699
a6cf197c1f1e also try 'metis' with 'full_types'
blanchet
parents: 57697
diff changeset
    53
  val infer_formulas_types : Proof.context -> term list -> term list
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
    54
  val introduce_spass_skolem : (term, string) atp_step list -> (term, string) atp_step list
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
    55
  val factify_atp_proof : (string * 'a) list -> term list -> term -> (term, string) atp_step list ->
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
    56
    (term, string) atp_step list
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    57
end;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    58
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45887
diff changeset
    59
structure ATP_Proof_Reconstruct : ATP_PROOF_RECONSTRUCT =
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    60
struct
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    61
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43062
diff changeset
    62
open ATP_Util
38028
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
    63
open ATP_Problem
39452
70a57e40f795 factored out TSTP/SPASS/Vampire proof parsing;
blanchet
parents: 39425
diff changeset
    64
open ATP_Proof
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45887
diff changeset
    65
open ATP_Problem_Generate
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    66
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    67
val metisN = "metis"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    68
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    69
val full_typesN = "full_types"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    70
val partial_typesN = "partial_types"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    71
val no_typesN = "no_types"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    72
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    73
val really_full_type_enc = "mono_tags"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    74
val full_type_enc = "poly_guards_query"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    75
val partial_type_enc = "poly_args"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    76
val no_type_enc = "erased"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    77
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    78
val full_type_encs = [full_type_enc, really_full_type_enc]
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    79
val partial_type_encs = partial_type_enc :: full_type_encs
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    80
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    81
val type_enc_aliases =
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    82
  [(full_typesN, full_type_encs),
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    83
   (partial_typesN, partial_type_encs),
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    84
   (no_typesN, [no_type_enc])]
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    85
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    86
fun unalias_type_enc s =
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    87
  AList.lookup (op =) type_enc_aliases s |> the_default [s]
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    88
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    89
val default_metis_lam_trans = combsN
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45511
diff changeset
    90
50670
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 49983
diff changeset
    91
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
    92
  | term_name' _ = ""
50670
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 49983
diff changeset
    93
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 49983
diff changeset
    94
fun lambda' v = Term.lambda_name (term_name' v, v)
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 49983
diff changeset
    95
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 49983
diff changeset
    96
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
    97
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
    98
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
    99
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
   100
  let val ww = "'" ^ w in
56254
a2dd9200854d more antiquotations;
wenzelm
parents: 56104
diff changeset
   101
    TFree (ww, the_default @{sort type} (Variable.def_sort ctxt (ww, ~1)))
43135
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   102
  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
   103
54820
d9ab86c044fd extended ATP types with sorts
blanchet
parents: 54819
diff changeset
   104
exception ATP_CLASS of string list
54818
a80bd631e573 honor SPASS-Pirate type arguments
blanchet
parents: 54811
diff changeset
   105
exception ATP_TYPE of string atp_type list
54811
df56a01f5684 parse SPASS-Pirate types
blanchet
parents: 54799
diff changeset
   106
exception ATP_TERM of (string, string atp_type) atp_term list
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52758
diff changeset
   107
exception ATP_FORMULA of
54811
df56a01f5684 parse SPASS-Pirate types
blanchet
parents: 54799
diff changeset
   108
  (string, string, (string, string atp_type) atp_term, string) atp_formula list
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   109
exception SAME of unit
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   110
54820
d9ab86c044fd extended ATP types with sorts
blanchet
parents: 54819
diff changeset
   111
fun class_of_atp_class cls =
d9ab86c044fd extended ATP types with sorts
blanchet
parents: 54819
diff changeset
   112
  (case unprefix_and_unascii class_prefix cls of
d9ab86c044fd extended ATP types with sorts
blanchet
parents: 54819
diff changeset
   113
    SOME s => s
d9ab86c044fd extended ATP types with sorts
blanchet
parents: 54819
diff changeset
   114
  | NONE => raise ATP_CLASS [cls])
d9ab86c044fd extended ATP types with sorts
blanchet
parents: 54819
diff changeset
   115
54811
df56a01f5684 parse SPASS-Pirate types
blanchet
parents: 54799
diff changeset
   116
(* Type variables are given the basic sort "HOL.type". Some will later be constrained by information
df56a01f5684 parse SPASS-Pirate types
blanchet
parents: 54799
diff changeset
   117
   from type literals, or by type inference. *)
54820
d9ab86c044fd extended ATP types with sorts
blanchet
parents: 54819
diff changeset
   118
fun typ_of_atp_type ctxt (ty as AType ((a, clss), tys)) =
57697
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   119
    let val Ts = map (typ_of_atp_type ctxt) tys in
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   120
      (case unprefix_and_unascii type_const_prefix a of
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   121
        SOME b => Type (invert_const b, Ts)
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   122
      | NONE =>
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   123
        if not (null tys) then
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   124
          raise ATP_TYPE [ty] (* only "tconst"s have type arguments *)
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   125
        else
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   126
          (case unprefix_and_unascii tfree_prefix a of
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   127
            SOME b => make_tfree ctxt b
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   128
          | NONE =>
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   129
            (* The term could be an Isabelle variable or a variable from the ATP, say "X1" or "_5018".
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   130
               Sometimes variables from the ATP are indistinguishable from Isabelle variables, which
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   131
               forces us to use a type parameter in all cases. *)
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   132
            Type_Infer.param 0 ("'" ^ perhaps (unprefix_and_unascii tvar_prefix) a,
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   133
              (if null clss then @{sort type} else map class_of_atp_class clss))))
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   134
    end
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   135
  | typ_of_atp_type ctxt (AFun (ty1, ty2)) = typ_of_atp_type ctxt ty1 --> typ_of_atp_type ctxt ty2
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   136
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   137
fun atp_type_of_atp_term (ATerm ((s, _), us)) =
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   138
  let val tys = map atp_type_of_atp_term us in
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   139
    if s = tptp_fun_type then
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   140
      (case tys of
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   141
        [ty1, ty2] => AFun (ty1, ty2)
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   142
      | _ => raise ATP_TYPE tys)
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   143
    else
44341963ade3 correctly translate THF functions from terms to types
blanchet
parents: 57635
diff changeset
   144
      AType ((s, []), tys)
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   145
  end
54818
a80bd631e573 honor SPASS-Pirate type arguments
blanchet
parents: 54811
diff changeset
   146
a80bd631e573 honor SPASS-Pirate type arguments
blanchet
parents: 54811
diff changeset
   147
fun typ_of_atp_term ctxt = typ_of_atp_type ctxt o atp_type_of_atp_term
a80bd631e573 honor SPASS-Pirate type arguments
blanchet
parents: 54811
diff changeset
   148
54789
blanchet
parents: 54788
diff changeset
   149
(* Type class literal applied to a type. Returns triple of polarity, class, type. *)
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   150
fun type_constraint_of_term ctxt (u as ATerm ((a, _), us)) =
54818
a80bd631e573 honor SPASS-Pirate type arguments
blanchet
parents: 54811
diff changeset
   151
  (case (unprefix_and_unascii class_prefix a, map (typ_of_atp_term ctxt) us) of
42606
0c76cf483899 show sorts not just types in Isar proofs + tuning
blanchet
parents: 42595
diff changeset
   152
    (SOME b, [T]) => (b, T)
54789
blanchet
parents: 54788
diff changeset
   153
  | _ => raise ATP_TERM [u])
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   154
43907
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   155
(* Accumulate type constraints in a formula: negative type literals. *)
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   156
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
   157
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
   158
  | 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
   159
  | add_type_constraint _ _ = I
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   160
55185
a0bd8d6414e6 centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents: 54822
diff changeset
   161
fun repair_var_name_raw s =
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   162
  let
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   163
    fun subscript_name s n = s ^ nat_subscript n
50704
cd1fcda1ea88 rename variable in binder, not just in body
blanchet
parents: 50703
diff changeset
   164
    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
   165
  in
54789
blanchet
parents: 54788
diff changeset
   166
    (case space_explode "_" s of
blanchet
parents: 54788
diff changeset
   167
      [_] =>
blanchet
parents: 54788
diff changeset
   168
      (case take_suffix Char.isDigit (String.explode s) of
blanchet
parents: 54788
diff changeset
   169
        (cs1 as _ :: _, cs2 as _ :: _) =>
blanchet
parents: 54788
diff changeset
   170
        subscript_name (String.implode cs1) (the (Int.fromString (String.implode cs2)))
blanchet
parents: 54788
diff changeset
   171
      | (_, _) => s)
blanchet
parents: 54788
diff changeset
   172
    | [s1, s2] => (case Int.fromString s2 of SOME n => subscript_name s1 n | NONE => s)
blanchet
parents: 54788
diff changeset
   173
    | _ => s)
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   174
  end
43182
649bada59658 slacker version of "fastype_of", in case a function has dummy type
blanchet
parents: 43176
diff changeset
   175
55185
a0bd8d6414e6 centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents: 54822
diff changeset
   176
fun repair_var_name textual s =
a0bd8d6414e6 centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents: 54822
diff changeset
   177
  (case unprefix_and_unascii schematic_var_prefix s of
a0bd8d6414e6 centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents: 54822
diff changeset
   178
    SOME s => s
a0bd8d6414e6 centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents: 54822
diff changeset
   179
  | NONE => s |> textual ? repair_var_name_raw)
a0bd8d6414e6 centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents: 54822
diff changeset
   180
54789
blanchet
parents: 54788
diff changeset
   181
(* The number of type arguments of a constant, zero if it's monomorphic. For (instances of) Skolem
blanchet
parents: 54788
diff changeset
   182
   pseudoconstants, this information is encoded in the constant name. *)
52125
ac7830871177 improved handling of free variables' types in Isar proofs
blanchet
parents: 52034
diff changeset
   183
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
   184
  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
   185
    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
   186
  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
   187
    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
   188
  else
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   189
    (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
   190
56254
a2dd9200854d more antiquotations;
wenzelm
parents: 56104
diff changeset
   191
fun slack_fastype_of t = fastype_of t handle TERM _ => Type_Infer.anyT @{sort type}
43182
649bada59658 slacker version of "fastype_of", in case a function has dummy type
blanchet
parents: 43176
diff changeset
   192
51192
4dc6bb65c3c3 slacker comparison for Skolems, to avoid trivial equations
blanchet
parents: 50704
diff changeset
   193
(* 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
   194
fun loose_aconv (Free (s, _), Free (s', _)) = s = s'
4dc6bb65c3c3 slacker comparison for Skolems, to avoid trivial equations
blanchet
parents: 50704
diff changeset
   195
  | loose_aconv (t, t') = t aconv t'
4dc6bb65c3c3 slacker comparison for Skolems, to avoid trivial equations
blanchet
parents: 50704
diff changeset
   196
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   197
val spass_skolem_prefix = "sk" (* "skc" or "skf" *)
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   198
val vampire_skolem_prefix = "sK"
50703
76a2e506c125 swap Vampire's Skolem arguments to bring them in line with what E and metis's new skolemizer do (helps Isar proof reconstruction in some cases)
blanchet
parents: 50693
diff changeset
   199
57257
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   200
fun var_index_of_textual textual = if textual then 0 else 1
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   201
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   202
fun quantify_over_var textual quant_of var_s var_T t =
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   203
  let
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   204
    val vars = ((var_s, var_index_of_textual textual), var_T) ::
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   205
                 filter (fn ((s, _), _) => s = var_s) (Term.add_vars t [])
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   206
    val normTs = vars |> AList.group (op =) |> map (apsnd hd)
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   207
    fun norm_var_types (Var (x, T)) =
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   208
        Var (x, the_default T (AList.lookup (op =) normTs x))
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   209
      | norm_var_types t = t
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   210
  in t |> map_aterms norm_var_types |> fold_rev quant_of (map Var normTs) end
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   211
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   212
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   213
(* Higher-order translation. Variables are typed (although we don't use that information). Lambdas
57256
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   214
   are typed.
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   215
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   216
  The code is similar to term_of_atp_fo. *)
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   217
fun term_of_atp_ho ctxt textual sym_tab =
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   218
  let
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   219
    val thy = Proof_Context.theory_of ctxt
57257
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   220
    val var_index = var_index_of_textual textual
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   221
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   222
    fun do_term opt_T u =
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   223
      (case u of
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   224
        AAbs(((var, ty), term), []) =>
57257
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   225
        let
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   226
          val typ = typ_of_atp_type ctxt ty
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   227
          val var_name = repair_var_name textual var
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   228
          val tm = do_term NONE term
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   229
        in quantify_over_var textual lambda' var_name typ tm end
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   230
      | ATerm ((s, tys), us) =>
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   231
        if s = ""
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   232
        then error "Isar proof reconstruction failed because the ATP proof \
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   233
                     \contains unparsable material."
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   234
        else if s = tptp_equal then
57703
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   235
          let
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   236
            val ts = map (do_term NONE) us
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   237
            fun equal_term ts =
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   238
                list_comb (Const (@{const_name HOL.eq}, Type_Infer.anyT @{sort type}), ts)
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   239
          in
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   240
            if textual then
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   241
              (case ts of
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   242
                [fst, lst] =>
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   243
                if loose_aconv (fst, lst) then
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   244
                  @{const True}
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   245
                else if Term.aconv_untyped (lst, @{const True}) then
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   246
                  fst
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   247
                else if Term.aconv_untyped (fst, @{const True}) then
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   248
                  lst
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   249
                else
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   250
                  equal_term ts
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   251
              | _ => equal_term ts)
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   252
            else
57703
fefe3ea75289 removing the '= True' generated by Leo-II.
fleury
parents: 57699
diff changeset
   253
              equal_term ts
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   254
          end
57256
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   255
        else if not (null us) then
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   256
          let
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   257
            val args = List.map (do_term NONE) us
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   258
            val opt_T' = SOME (map slack_fastype_of args ---> the_default dummyT opt_T)
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   259
            val func = do_term opt_T' (ATerm ((s, tys), []))
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   260
          in foldl1 (op $) (func :: args) end
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   261
        else if s = tptp_or then HOLogic.disj
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   262
        else if s = tptp_and then HOLogic.conj
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   263
        else if s = tptp_implies then HOLogic.imp
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   264
        else if s = tptp_iff orelse s = tptp_equal then HOLogic.eq_const dummyT
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   265
        else if s = tptp_not_iff orelse s = tptp_not_equal then @{term "% P Q. Q ~= P"}
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   266
        else if s = tptp_if then @{term "% P Q. Q --> P"}
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   267
        else if s = tptp_not_and then @{term "% P Q. ~ (P & Q)"}
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   268
        else if s = tptp_not_or then @{term "% P Q. ~ (P | Q)"}
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   269
        else if s = tptp_not then HOLogic.Not
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   270
        else if s = tptp_ho_forall then HOLogic.all_const dummyT
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   271
        else if s = tptp_ho_exists then HOLogic.exists_const dummyT
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   272
        else
57256
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   273
          (case unprefix_and_unascii const_prefix s of
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   274
            SOME s' =>
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   275
            let
57547
677b07d777c3 don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
blanchet
parents: 57267
diff changeset
   276
              val ((s', _), mangled_us) = s' |> unmangled_const |>> `invert_const
57256
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   277
              val num_ty_args = length us - the_default 0 (Symtab.lookup sym_tab s)
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   278
              val (type_us, term_us) = chop num_ty_args us |>> append mangled_us
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   279
              val term_ts = map (do_term NONE) term_us
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   280
              val Ts = map (typ_of_atp_type ctxt) tys @ map (typ_of_atp_term ctxt) type_us
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   281
              val T =
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   282
                (if not (null Ts) andalso robust_const_num_type_args thy s' = length Ts then
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   283
                   if textual then try (Sign.const_instance thy) (s', Ts) else NONE
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   284
                 else
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   285
                   NONE)
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   286
                |> (fn SOME T => T
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   287
                     | NONE =>
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   288
                       map slack_fastype_of term_ts --->
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   289
                       the_default (Type_Infer.anyT @{sort type}) opt_T)
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   290
              val t = Const (unproxify_const s', T)
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   291
            in list_comb (t, term_ts) end
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   292
          | NONE => (* a free or schematic variable *)
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   293
            let
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   294
              fun fresh_up s =
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   295
                [(s, ())] |> Variable.variant_frees ctxt [] |> hd |> fst
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   296
              val ts = map (do_term NONE) us
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   297
              val T =
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   298
                (case opt_T of
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   299
                  SOME T => map slack_fastype_of ts ---> T
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   300
                | NONE =>
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   301
                  map slack_fastype_of ts --->
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   302
                    (case tys of
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   303
                      [ty] => typ_of_atp_type ctxt ty
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   304
                    | _ => Type_Infer.anyT @{sort type}))
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   305
              val t =
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   306
                (case unprefix_and_unascii fixed_var_prefix s of
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   307
                  SOME s => Free (s, T)
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   308
                | NONE =>
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   309
                  if textual andalso not (is_tptp_variable s) then
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   310
                    Free (s |> textual ? (repair_var_name_raw #> fresh_up), T)
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   311
                  else
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   312
                    Var ((repair_var_name textual s, var_index), T))
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   313
            in list_comb (t, ts) end))
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   314
  in do_term end
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   315
56254
a2dd9200854d more antiquotations;
wenzelm
parents: 56104
diff changeset
   316
(* First-order translation. No types are known for variables. "Type_Infer.anyT @{sort type}"
a2dd9200854d more antiquotations;
wenzelm
parents: 56104
diff changeset
   317
   should allow them to be inferred. *)
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   318
fun term_of_atp_fo ctxt textual sym_tab =
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   319
  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
   320
    val thy = Proof_Context.theory_of ctxt
54789
blanchet
parents: 54788
diff changeset
   321
    (* For Metis, we use 1 rather than 0 because variable references in clauses may otherwise
blanchet
parents: 54788
diff changeset
   322
       conflict with variable constraints in the goal. At least, type inference often fails
blanchet
parents: 54788
diff changeset
   323
       otherwise. See also "axiom_inference" in "Metis_Reconstruct". *)
57257
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   324
    val var_index = var_index_of_textual textual
57199
blanchet
parents: 56854
diff changeset
   325
43131
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   326
    fun do_term extra_ts opt_T u =
54789
blanchet
parents: 54788
diff changeset
   327
      (case u of
54818
a80bd631e573 honor SPASS-Pirate type arguments
blanchet
parents: 54811
diff changeset
   328
        ATerm ((s, tys), us) =>
57199
blanchet
parents: 56854
diff changeset
   329
        if s = "" then
blanchet
parents: 56854
diff changeset
   330
          error "Isar proof reconstruction failed because the ATP proof contains unparsable \
blanchet
parents: 56854
diff changeset
   331
            \material."
51878
f11039b31bae handle dummy atp terms
smolkas
parents: 51192
diff changeset
   332
        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
   333
          @{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
   334
        else if s = tptp_equal then
43093
blanchet
parents: 43085
diff changeset
   335
          let val ts = map (do_term [] NONE) us in
54822
d4a56c826769 pick up tfree/tvar type from SPASS-Pirate proof
blanchet
parents: 54821
diff changeset
   336
            if textual andalso length ts = 2 andalso loose_aconv (hd ts, List.last ts) then
39106
5ab6a3707499 fix trivial "x = x" fact detection
blanchet
parents: 38988
diff changeset
   337
              @{const True}
5ab6a3707499 fix trivial "x = x" fact detection
blanchet
parents: 38988
diff changeset
   338
            else
56254
a2dd9200854d more antiquotations;
wenzelm
parents: 56104
diff changeset
   339
              list_comb (Const (@{const_name HOL.eq}, Type_Infer.anyT @{sort type}), ts)
39106
5ab6a3707499 fix trivial "x = x" fact detection
blanchet
parents: 38988
diff changeset
   340
          end
54789
blanchet
parents: 54788
diff changeset
   341
        else
blanchet
parents: 54788
diff changeset
   342
          (case unprefix_and_unascii const_prefix s of
blanchet
parents: 54788
diff changeset
   343
            SOME s' =>
54818
a80bd631e573 honor SPASS-Pirate type arguments
blanchet
parents: 54811
diff changeset
   344
            let val ((s', s''), mangled_us) = s' |> unmangled_const |>> `invert_const in
54789
blanchet
parents: 54788
diff changeset
   345
              if s' = type_tag_name then
blanchet
parents: 54788
diff changeset
   346
                (case mangled_us @ us of
54818
a80bd631e573 honor SPASS-Pirate type arguments
blanchet
parents: 54811
diff changeset
   347
                  [typ_u, term_u] => do_term extra_ts (SOME (typ_of_atp_term ctxt typ_u)) term_u
54789
blanchet
parents: 54788
diff changeset
   348
                | _ => raise ATP_TERM us)
blanchet
parents: 54788
diff changeset
   349
              else if s' = predicator_name then
blanchet
parents: 54788
diff changeset
   350
                do_term [] (SOME @{typ bool}) (hd us)
blanchet
parents: 54788
diff changeset
   351
              else if s' = app_op_name then
blanchet
parents: 54788
diff changeset
   352
                let val extra_t = do_term [] NONE (List.last us) in
blanchet
parents: 54788
diff changeset
   353
                  do_term (extra_t :: extra_ts)
57199
blanchet
parents: 56854
diff changeset
   354
                    (case opt_T of SOME T => SOME (slack_fastype_of extra_t --> T) | NONE => NONE)
blanchet
parents: 56854
diff changeset
   355
                    (nth us (length us - 2))
54789
blanchet
parents: 54788
diff changeset
   356
                end
blanchet
parents: 54788
diff changeset
   357
              else if s' = type_guard_name then
blanchet
parents: 54788
diff changeset
   358
                @{const True} (* ignore type predicates *)
blanchet
parents: 54788
diff changeset
   359
              else
blanchet
parents: 54788
diff changeset
   360
                let
blanchet
parents: 54788
diff changeset
   361
                  val new_skolem = String.isPrefix new_skolem_const_prefix s''
54818
a80bd631e573 honor SPASS-Pirate type arguments
blanchet
parents: 54811
diff changeset
   362
                  val num_ty_args = length us - the_default 0 (Symtab.lookup sym_tab s)
a80bd631e573 honor SPASS-Pirate type arguments
blanchet
parents: 54811
diff changeset
   363
                  val (type_us, term_us) = chop num_ty_args us |>> append mangled_us
54789
blanchet
parents: 54788
diff changeset
   364
                  val term_ts = map (do_term [] NONE) term_us
54818
a80bd631e573 honor SPASS-Pirate type arguments
blanchet
parents: 54811
diff changeset
   365
a80bd631e573 honor SPASS-Pirate type arguments
blanchet
parents: 54811
diff changeset
   366
                  val Ts = map (typ_of_atp_type ctxt) tys @ map (typ_of_atp_term ctxt) type_us
54789
blanchet
parents: 54788
diff changeset
   367
                  val T =
54818
a80bd631e573 honor SPASS-Pirate type arguments
blanchet
parents: 54811
diff changeset
   368
                    (if not (null Ts) andalso robust_const_num_type_args thy s' = length Ts then
57199
blanchet
parents: 56854
diff changeset
   369
                       if new_skolem then SOME (Type_Infer.paramify_vars (tl Ts ---> hd Ts))
blanchet
parents: 56854
diff changeset
   370
                       else if textual then try (Sign.const_instance thy) (s', Ts)
blanchet
parents: 56854
diff changeset
   371
                       else NONE
54789
blanchet
parents: 54788
diff changeset
   372
                     else
blanchet
parents: 54788
diff changeset
   373
                       NONE)
blanchet
parents: 54788
diff changeset
   374
                    |> (fn SOME T => T
blanchet
parents: 54788
diff changeset
   375
                         | NONE =>
56254
a2dd9200854d more antiquotations;
wenzelm
parents: 56104
diff changeset
   376
                           map slack_fastype_of term_ts --->
57199
blanchet
parents: 56854
diff changeset
   377
                           the_default (Type_Infer.anyT @{sort type}) opt_T)
54789
blanchet
parents: 54788
diff changeset
   378
                  val t =
blanchet
parents: 54788
diff changeset
   379
                    if new_skolem then Var ((new_skolem_var_name_of_const s'', var_index), T)
blanchet
parents: 54788
diff changeset
   380
                    else Const (unproxify_const s', T)
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   381
                in
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   382
                  list_comb (t, term_ts @ extra_ts)
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   383
                end
54789
blanchet
parents: 54788
diff changeset
   384
            end
blanchet
parents: 54788
diff changeset
   385
          | NONE => (* a free or schematic variable *)
blanchet
parents: 54788
diff changeset
   386
            let
blanchet
parents: 54788
diff changeset
   387
              (* This assumes that distinct names are mapped to distinct names by
54822
d4a56c826769 pick up tfree/tvar type from SPASS-Pirate proof
blanchet
parents: 54821
diff changeset
   388
                 "Variable.variant_frees". This does not hold in general but should hold for
d4a56c826769 pick up tfree/tvar type from SPASS-Pirate proof
blanchet
parents: 54821
diff changeset
   389
                 ATP-generated Skolem function names, since these end with a digit and
d4a56c826769 pick up tfree/tvar type from SPASS-Pirate proof
blanchet
parents: 54821
diff changeset
   390
                 "variant_frees" appends letters. *)
57199
blanchet
parents: 56854
diff changeset
   391
              fun fresh_up s = [(s, ())] |> Variable.variant_frees ctxt [] |> hd |> fst
blanchet
parents: 56854
diff changeset
   392
54789
blanchet
parents: 54788
diff changeset
   393
              val term_ts =
blanchet
parents: 54788
diff changeset
   394
                map (do_term [] NONE) us
blanchet
parents: 54788
diff changeset
   395
                (* SPASS (3.8ds) and Vampire (2.6) pass arguments to Skolem functions in reverse
57699
a6cf197c1f1e also try 'metis' with 'full_types'
blanchet
parents: 57697
diff changeset
   396
                   order, which is incompatible with "metis"'s new skolemizer. *)
54789
blanchet
parents: 54788
diff changeset
   397
                |> exists (fn pre => String.isPrefix pre s)
blanchet
parents: 54788
diff changeset
   398
                  [spass_skolem_prefix, vampire_skolem_prefix] ? rev
blanchet
parents: 54788
diff changeset
   399
              val ts = term_ts @ extra_ts
blanchet
parents: 54788
diff changeset
   400
              val T =
blanchet
parents: 54788
diff changeset
   401
                (case opt_T of
blanchet
parents: 54788
diff changeset
   402
                  SOME T => map slack_fastype_of term_ts ---> T
54822
d4a56c826769 pick up tfree/tvar type from SPASS-Pirate proof
blanchet
parents: 54821
diff changeset
   403
                | NONE =>
d4a56c826769 pick up tfree/tvar type from SPASS-Pirate proof
blanchet
parents: 54821
diff changeset
   404
                  map slack_fastype_of ts --->
56254
a2dd9200854d more antiquotations;
wenzelm
parents: 56104
diff changeset
   405
                  (case tys of [ty] => typ_of_atp_type ctxt ty | _ => Type_Infer.anyT @{sort type}))
54789
blanchet
parents: 54788
diff changeset
   406
              val t =
blanchet
parents: 54788
diff changeset
   407
                (case unprefix_and_unascii fixed_var_prefix s of
blanchet
parents: 54788
diff changeset
   408
                  SOME s => Free (s, T)
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   409
                | NONE =>
55185
a0bd8d6414e6 centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents: 54822
diff changeset
   410
                  if textual andalso not (is_tptp_variable s) then
a0bd8d6414e6 centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents: 54822
diff changeset
   411
                    Free (s |> textual ? (repair_var_name_raw #> fresh_up), T)
a0bd8d6414e6 centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents: 54822
diff changeset
   412
                  else
a0bd8d6414e6 centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents: 54822
diff changeset
   413
                    Var ((repair_var_name textual s, var_index), T))
54789
blanchet
parents: 54788
diff changeset
   414
            in list_comb (t, ts) end))
43093
blanchet
parents: 43085
diff changeset
   415
  in do_term [] end
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   416
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   417
fun term_of_atp ctxt (ATP_Problem.THF _) type_enc =
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   418
    if ATP_Problem_Generate.is_type_enc_higher_order type_enc
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   419
    then term_of_atp_ho ctxt
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   420
    else error "Unsupported Isar reconstruction."
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   421
  | term_of_atp ctxt _ type_enc =
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   422
    if not (ATP_Problem_Generate.is_type_enc_higher_order type_enc)
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   423
    then term_of_atp_fo ctxt
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   424
    else error "Unsupported Isar reconstruction."
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   425
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   426
fun term_of_atom ctxt format type_enc textual sym_tab pos (u as ATerm ((s, _), _)) =
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   427
  if String.isPrefix class_prefix s then
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51878
diff changeset
   428
    add_type_constraint pos (type_constraint_of_term ctxt u)
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   429
    #> pair @{const True}
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   430
  else
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   431
    pair (term_of_atp ctxt format type_enc textual sym_tab (SOME @{typ bool}) u)
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   432
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   433
(* Update schematic type variables with detected sort constraints. It's not
42563
e70ffe3846d0 improve helper type instantiation code
blanchet
parents: 42562
diff changeset
   434
   totally clear whether this code is necessary. *)
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   435
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
   436
  let
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   437
    fun do_type (Type (a, Ts)) = Type (a, map do_type Ts)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   438
      | do_type (TVar (xi, s)) =
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   439
        TVar (xi, the_default s (Vartab.lookup tvar_tab xi))
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   440
      | do_type (TFree z) = TFree z
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   441
    fun do_term (Const (a, T)) = Const (a, do_type T)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   442
      | do_term (Free (a, T)) = Free (a, do_type T)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   443
      | do_term (Var (xi, T)) = Var (xi, do_type T)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   444
      | do_term (t as Bound _) = t
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   445
      | 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
   446
      | do_term (t1 $ t2) = do_term t1 $ do_term t2
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   447
  in t |> not (Vartab.is_empty tvar_tab) ? do_term end
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   448
54811
df56a01f5684 parse SPASS-Pirate types
blanchet
parents: 54799
diff changeset
   449
(* Interpret an ATP formula as a HOL term, extracting sort constraints as they appear in the
df56a01f5684 parse SPASS-Pirate types
blanchet
parents: 54799
diff changeset
   450
   formula. *)
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   451
fun prop_of_atp ctxt format type_enc textual sym_tab phi =
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   452
  let
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   453
    fun do_formula pos phi =
54789
blanchet
parents: 54788
diff changeset
   454
      (case phi of
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   455
        AQuant (_, [], phi) => do_formula pos phi
42526
46d485f8d144 added room for types in ATP quantifiers
blanchet
parents: 42451
diff changeset
   456
      | AQuant (q, (s, _) :: xs, phi') =>
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   457
        do_formula pos (AQuant (q, xs, phi'))
42526
46d485f8d144 added room for types in ATP quantifiers
blanchet
parents: 42451
diff changeset
   458
        (* FIXME: TFF *)
57257
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   459
        #>> quantify_over_var textual (case q of AForall => forall_of | AExists => exists_of)
0d5e34ba4d28 Correcting the type parser
fleury
parents: 57256
diff changeset
   460
          (repair_var_name textual s) dummyT
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   461
      | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   462
      | AConn (c, [phi1, phi2]) =>
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   463
        do_formula (pos |> c = AImplies ? not) phi1
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   464
        ##>> do_formula pos phi2
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   465
        #>> (case c of
54811
df56a01f5684 parse SPASS-Pirate types
blanchet
parents: 54799
diff changeset
   466
              AAnd => s_conj
df56a01f5684 parse SPASS-Pirate types
blanchet
parents: 54799
diff changeset
   467
            | AOr => s_disj
df56a01f5684 parse SPASS-Pirate types
blanchet
parents: 54799
diff changeset
   468
            | AImplies => s_imp
df56a01f5684 parse SPASS-Pirate types
blanchet
parents: 54799
diff changeset
   469
            | AIff => s_iff
df56a01f5684 parse SPASS-Pirate types
blanchet
parents: 54799
diff changeset
   470
            | ANot => raise Fail "impossible connective")
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   471
      | AAtom tm => term_of_atom ctxt format type_enc textual sym_tab pos tm
54789
blanchet
parents: 54788
diff changeset
   472
      | _ => raise ATP_FORMULA [phi])
blanchet
parents: 54788
diff changeset
   473
  in
blanchet
parents: 54788
diff changeset
   474
    repair_tvar_sorts (do_formula true phi Vartab.empty)
blanchet
parents: 54788
diff changeset
   475
  end
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   476
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   477
val unprefix_fact_number = space_implode "_" o tl o space_explode "_"
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   478
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   479
fun resolve_fact facts s =
54789
blanchet
parents: 54788
diff changeset
   480
  (case try (unprefix fact_prefix) s of
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   481
    SOME s' =>
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   482
    let val s' = s' |> unprefix_fact_number |> unascii_of in
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   483
      AList.lookup (op =) facts s' |> Option.map (pair s')
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   484
    end
54789
blanchet
parents: 54788
diff changeset
   485
  | NONE => NONE)
54506
blanchet
parents: 54505
diff changeset
   486
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   487
fun resolve_conjecture s =
54789
blanchet
parents: 54788
diff changeset
   488
  (case try (unprefix conjecture_prefix) s of
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   489
    SOME s' => Int.fromString s'
54789
blanchet
parents: 54788
diff changeset
   490
  | NONE => NONE)
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   491
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   492
fun resolve_facts facts = map_filter (resolve_fact facts)
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   493
val resolve_conjectures = map_filter resolve_conjecture
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   494
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   495
fun is_axiom_used_in_proof pred =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   496
  exists (fn ((_, ss), _, _, _, []) => exists pred ss | _ => false)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   497
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   498
val leo2_extcnf_equal_neg_rule = "extcnf_equal_neg"
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   499
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   500
fun add_fact ctxt facts ((_, ss), _, _, rule, deps) =
57547
677b07d777c3 don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
blanchet
parents: 57267
diff changeset
   501
  (if member (op =) [agsyhol_core_rule, leo2_extcnf_equal_neg_rule, satallax_core_rule] rule then
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 55285
diff changeset
   502
     insert (op =) (short_thm_name ctxt ext, (Global, General))
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   503
   else
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   504
     I)
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   505
  #> (if null deps then union (op =) (resolve_facts facts ss) else I)
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   506
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   507
fun used_facts_in_atp_proof ctxt facts atp_proof =
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   508
  if null atp_proof then facts else fold (add_fact ctxt facts) atp_proof []
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   509
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   510
fun used_facts_in_unsound_atp_proof _ _ [] = NONE
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   511
  | used_facts_in_unsound_atp_proof ctxt facts atp_proof =
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   512
    let val used_facts = used_facts_in_atp_proof ctxt facts atp_proof in
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   513
      if forall (fn (_, (sc, _)) => sc = Global) used_facts andalso
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   514
         not (is_axiom_used_in_proof (is_some o resolve_conjecture) atp_proof) then
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   515
        SOME (map fst used_facts)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   516
      else
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   517
        NONE
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   518
    end
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   519
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   520
val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   521
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   522
(* overapproximation (good enough) *)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   523
fun is_lam_lifted s =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   524
  String.isPrefix fact_prefix s andalso
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   525
  String.isSubstring ascii_of_lam_fact_prefix s
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   526
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   527
val is_combinator_def = String.isPrefix (helper_prefix ^ combinator_prefix)
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   528
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55195
diff changeset
   529
fun atp_proof_prefers_lifting atp_proof =
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55195
diff changeset
   530
  (is_axiom_used_in_proof is_combinator_def atp_proof,
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55195
diff changeset
   531
   is_axiom_used_in_proof is_lam_lifted atp_proof) = (false, true)
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   532
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   533
val is_typed_helper_name =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   534
  String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   535
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   536
fun is_typed_helper_used_in_atp_proof atp_proof =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   537
  is_axiom_used_in_proof is_typed_helper_name atp_proof
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   538
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   539
fun replace_one_dependency (old, new) dep = if is_same_atp_step dep old then new else [dep]
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   540
fun replace_dependencies_in_line old_new (name, role, t, rule, deps) =
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   541
  (name, role, t, rule, fold (union (op =) o replace_one_dependency old_new) deps [])
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   542
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   543
fun repair_name "$true" = "c_True"
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   544
  | repair_name "$false" = "c_False"
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   545
  | repair_name "$$e" = tptp_equal (* seen in Vampire proofs *)
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   546
  | repair_name s =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   547
    if is_tptp_equal s orelse
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   548
       (* seen in Vampire proofs *)
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   549
       (String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   550
      tptp_equal
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   551
    else
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   552
      s
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   553
57699
a6cf197c1f1e also try 'metis' with 'full_types'
blanchet
parents: 57697
diff changeset
   554
fun set_var_index j = map_aterms (fn Var ((s, 0), T) => Var ((s, j), T) | t => t)
57635
97adb86619a4 more robust handling of types for skolems (modeled as Frees)
blanchet
parents: 57547
diff changeset
   555
57699
a6cf197c1f1e also try 'metis' with 'full_types'
blanchet
parents: 57697
diff changeset
   556
fun infer_formulas_types ctxt =
57635
97adb86619a4 more robust handling of types for skolems (modeled as Frees)
blanchet
parents: 57547
diff changeset
   557
  map_index (uncurry (fn j => set_var_index j #> Type.constraint HOLogic.boolT))
97adb86619a4 more robust handling of types for skolems (modeled as Frees)
blanchet
parents: 57547
diff changeset
   558
  #> Syntax.check_terms (Proof_Context.set_mode Proof_Context.mode_schematic ctxt)
97adb86619a4 more robust handling of types for skolems (modeled as Frees)
blanchet
parents: 57547
diff changeset
   559
  #> map (set_var_index 0)
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   560
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   561
val combinator_table =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   562
  [(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def [abs_def]}),
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   563
   (@{const_name Meson.COMBK}, @{thm Meson.COMBK_def [abs_def]}),
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   564
   (@{const_name Meson.COMBB}, @{thm Meson.COMBB_def [abs_def]}),
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   565
   (@{const_name Meson.COMBC}, @{thm Meson.COMBC_def [abs_def]}),
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   566
   (@{const_name Meson.COMBS}, @{thm Meson.COMBS_def [abs_def]})]
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   567
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   568
fun uncombine_term thy =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   569
  let
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   570
    fun aux (t1 $ t2) = betapply (pairself aux (t1, t2))
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   571
      | aux (Abs (s, T, t')) = Abs (s, T, aux t')
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   572
      | aux (t as Const (x as (s, _))) =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   573
        (case AList.lookup (op =) combinator_table s of
54756
blanchet
parents: 54507
diff changeset
   574
          SOME thm => thm |> prop_of |> specialize_type thy x |> Logic.dest_equals |> snd
blanchet
parents: 54507
diff changeset
   575
        | NONE => t)
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   576
      | aux t = t
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   577
  in aux end
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   578
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   579
fun unlift_term lifted =
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   580
  map_aterms (fn t as Const (s, _) =>
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   581
                 if String.isPrefix lam_lifted_prefix s then
54756
blanchet
parents: 54507
diff changeset
   582
                   (* FIXME: do something about the types *)
blanchet
parents: 54507
diff changeset
   583
                   (case AList.lookup (op =) lifted s of
blanchet
parents: 54507
diff changeset
   584
                     SOME t => unlift_term lifted t
blanchet
parents: 54507
diff changeset
   585
                   | NONE => t)
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   586
                 else
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   587
                   t
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   588
               | t => t)
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   589
57256
cf43583f9bb9 imported patch leo2_skolem_simplication
fleury
parents: 57255
diff changeset
   590
fun termify_line _ _ _ _ _ (_, Type_Role, _, _, _) = NONE
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   591
  | termify_line ctxt format type_enc lifted sym_tab (name, role, u, rule, deps) =
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   592
    let
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   593
      val thy = Proof_Context.theory_of ctxt
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   594
      val t = u
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   595
        |> prop_of_atp ctxt format type_enc true sym_tab
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   596
        |> uncombine_term thy
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   597
        |> unlift_term lifted
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   598
    in
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   599
      SOME (name, role, t, rule, deps)
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   600
    end
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   601
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   602
val waldmeister_conjecture_num = "1.0.0.0"
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   603
54756
blanchet
parents: 54507
diff changeset
   604
fun repair_waldmeister_endgame proof =
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   605
  let
56854
ddd3af5a683d fixed Waldmeister endgame w.r.t. "Trueprop"
blanchet
parents: 56254
diff changeset
   606
    fun repair_tail (name, _, @{const Trueprop} $ t, rule, deps) =
ddd3af5a683d fixed Waldmeister endgame w.r.t. "Trueprop"
blanchet
parents: 56254
diff changeset
   607
      (name, Negated_Conjecture, @{const Trueprop} $ s_not t, rule, deps)
54756
blanchet
parents: 54507
diff changeset
   608
    fun repair_body [] = []
blanchet
parents: 54507
diff changeset
   609
      | repair_body ((line as ((num, _), _, _, _, _)) :: lines) =
blanchet
parents: 54507
diff changeset
   610
        if num = waldmeister_conjecture_num then map repair_tail (line :: lines)
blanchet
parents: 54507
diff changeset
   611
        else line :: repair_body lines
blanchet
parents: 54507
diff changeset
   612
  in
blanchet
parents: 54507
diff changeset
   613
    repair_body proof
blanchet
parents: 54507
diff changeset
   614
  end
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   615
57635
97adb86619a4 more robust handling of types for skolems (modeled as Frees)
blanchet
parents: 57547
diff changeset
   616
fun map_proof_terms f (lines : ('a * 'b * 'c * 'd * 'e) list) =
97adb86619a4 more robust handling of types for skolems (modeled as Frees)
blanchet
parents: 57547
diff changeset
   617
  map2 (fn c => fn (a, b, _, d, e) => (a, b, c, d, e)) (f (map #3 lines)) lines
97adb86619a4 more robust handling of types for skolems (modeled as Frees)
blanchet
parents: 57547
diff changeset
   618
57258
67d85a8aa6cc Moving the remote prefix deleting on Sledgehammer's side
fleury
parents: 57257
diff changeset
   619
fun termify_atp_proof ctxt local_prover format type_enc pool lifted sym_tab =
57267
8b87114357bd integrated new Waldmeister code with 'sledgehammer' command
blanchet
parents: 57263
diff changeset
   620
  nasty_atp_proof pool
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   621
  #> map_term_names_in_atp_proof repair_name
57255
488046fdda59 add support for Isar reconstruction for thf1 ATP provers like Leo-II.
fleury
parents: 57199
diff changeset
   622
  #> map_filter (termify_line ctxt format type_enc lifted sym_tab)
57699
a6cf197c1f1e also try 'metis' with 'full_types'
blanchet
parents: 57697
diff changeset
   623
  #> map_proof_terms (infer_formulas_types ctxt #> map HOLogic.mk_Trueprop)
57258
67d85a8aa6cc Moving the remote prefix deleting on Sledgehammer's side
fleury
parents: 57257
diff changeset
   624
  #> local_prover = waldmeisterN ? repair_waldmeister_endgame
54499
319f8659267d refactoring
blanchet
parents: 53586
diff changeset
   625
55192
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   626
fun take_distinct_vars seen ((t as Var _) :: ts) =
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   627
    if member (op aconv) seen t then rev seen else take_distinct_vars (t :: seen) ts
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   628
  | take_distinct_vars seen _ = rev seen
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   629
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   630
fun unskolemize_term skos t =
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   631
  let
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   632
    val is_sko = member (op =) skos
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   633
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   634
    fun find_args args (t $ u) = find_args (u :: args) t #> find_args [] u
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   635
      | find_args _ (Abs (_, _, t)) = find_args [] t
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   636
      | find_args args (Free (s, _)) =
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   637
        if is_sko s then
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   638
          let val new = take_distinct_vars [] args in
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   639
            (fn [] => new | old => if length new < length old then new else old)
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   640
          end
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   641
        else
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   642
          I
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   643
      | find_args _ _ = I
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   644
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   645
    val alls = find_args [] t []
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   646
    val num_alls = length alls
55195
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   647
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   648
    fun fix_skos args (t $ u) = fix_skos (fix_skos [] u :: args) t
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   649
      | fix_skos args (t as Free (s, T)) =
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   650
        if is_sko s then list_comb (Free (s, funpow num_alls body_type T), drop num_alls args)
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   651
        else list_comb (t, args)
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   652
      | fix_skos [] (Abs (s, T, t)) = Abs (s, T, fix_skos [] t)
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   653
      | fix_skos [] t = t
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   654
      | fix_skos args t = list_comb (fix_skos [] t, args)
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   655
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   656
    val t' = fix_skos [] t
55192
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   657
55195
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   658
    fun add_sko (t as Free (s, _)) = is_sko s ? insert (op aconv) t
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   659
      | add_sko _ = I
55192
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   660
55195
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   661
    val exs = Term.fold_aterms add_sko t' []
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   662
  in
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   663
    t'
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   664
    |> HOLogic.dest_Trueprop
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   665
    |> fold exists_of exs |> Term.map_abs_vars (K Name.uu)
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   666
    |> fold_rev forall_of alls
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   667
    |> HOLogic.mk_Trueprop
55192
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   668
  end
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   669
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   670
fun introduce_spass_skolem [] = []
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   671
  | introduce_spass_skolem (proof as (_, _, _, rule1, _) :: _) =
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   672
    if rule1 = spass_input_rule then
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   673
      let
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   674
        fun add_sko (Free (s, _)) = String.isPrefix spass_skolem_prefix s ? insert (op =) s
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   675
          | add_sko _ = I
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   676
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   677
        (* union-find would be faster *)
54799
blanchet
parents: 54789
diff changeset
   678
        fun add_cycle [] = I
blanchet
parents: 54789
diff changeset
   679
          | add_cycle ss =
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   680
            fold (fn s => Graph.default_node (s, ())) ss
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   681
            #> fold Graph.add_edge (ss ~~ tl ss @ [hd ss])
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   682
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   683
        val (input_steps, other_steps) = List.partition (null o #5) proof
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   684
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   685
        val skoss = map (fn (_, _, t, _, _) => Term.fold_aterms add_sko t []) input_steps
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   686
        val skoss_input_steps = filter_out (null o fst) (skoss ~~ input_steps)
54799
blanchet
parents: 54789
diff changeset
   687
        val groups = Graph.strong_conn (fold add_cycle skoss Graph.empty)
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   688
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   689
        fun step_name_of_group skos = (implode skos, [])
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   690
        fun in_group group = member (op =) group o hd
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   691
        fun group_of sko = the (find_first (fn group => in_group group sko) groups)
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   692
55192
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   693
        fun new_steps (skoss_steps : (string list * (term, 'a) atp_step) list) group =
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   694
          let
55192
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   695
            val name = step_name_of_group group
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   696
            val name0 = name |>> prefix "0"
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   697
            val t =
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   698
              skoss_steps
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   699
              |> map (snd #> #3 #> HOLogic.dest_Trueprop)
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   700
              |> Library.foldr1 s_conj
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   701
              |> HOLogic.mk_Trueprop
55192
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   702
            val skos = fold (union (op =) o fst) skoss_steps []
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   703
            val deps = map (snd #> #1) skoss_steps
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   704
          in
55195
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   705
            [(name0, Unknown, unskolemize_term skos t, spass_pre_skolemize_rule, deps),
067142c53c3b reverted unsound optimization
blanchet
parents: 55192
diff changeset
   706
             (name, Unknown, t, spass_skolemize_rule, [name0])]
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   707
          end
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   708
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   709
        val sko_steps =
55192
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   710
          maps (fn group => new_steps (filter (in_group group o fst) skoss_input_steps) group)
b75b52c7cf94 unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
blanchet
parents: 55185
diff changeset
   711
            groups
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   712
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   713
        val old_news =
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   714
          map (fn (skos, (name, _, _, _, _)) => (name, [step_name_of_group (group_of skos)]))
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   715
            skoss_input_steps
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   716
        val repair_deps = fold replace_dependencies_in_line old_news
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   717
      in
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   718
        input_steps @ sko_steps @ map repair_deps other_steps
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   719
      end
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   720
  else
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   721
    proof
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   722
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   723
fun factify_atp_proof facts hyp_ts concl_t atp_proof =
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   724
  let
54799
blanchet
parents: 54789
diff changeset
   725
    fun factify_step ((num, ss), _, t, rule, deps) =
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   726
      let
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   727
        val (ss', role', t') =
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   728
          (case resolve_conjectures ss of
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   729
            [j] =>
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   730
            if j = length hyp_ts then ([], Conjecture, concl_t) else ([], Hypothesis, nth hyp_ts j)
55185
a0bd8d6414e6 centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents: 54822
diff changeset
   731
          | _ =>
57263
2b6a96cc64c9 simplified code
blanchet
parents: 57258
diff changeset
   732
            (case resolve_facts facts ss of
55185
a0bd8d6414e6 centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents: 54822
diff changeset
   733
              [] => (ss, Plain, t)
a0bd8d6414e6 centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
blanchet
parents: 54822
diff changeset
   734
            | facts => (map fst facts, Axiom, t)))
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   735
      in
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   736
        ((num, ss'), role', t', rule, deps)
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   737
      end
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   738
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   739
    val atp_proof = map factify_step atp_proof
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   740
    val names = map #1 atp_proof
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   741
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   742
    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
   743
    fun repair_deps (name, role, t, rule, deps) = (name, role, t, rule, map repair_dep deps)
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   744
  in
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   745
    map repair_deps atp_proof
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54770
diff changeset
   746
  end
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54501
diff changeset
   747
31038
immler@in.tum.de
parents: 31037
diff changeset
   748
end;