src/HOL/Tools/ATP/atp_reconstruct.ML
author blanchet
Sun, 06 Nov 2011 22:18:54 +0100
changeset 45378 67ed44d7c929
parent 45209 0e5e56e32bc0
child 45379 0147a4348ca1
permissions -rw-r--r--
more detailed preplay output
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43062
diff changeset
     1
(*  Title:      HOL/Tools/ATP/atp_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
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43062
diff changeset
     6
Proof reconstruction from ATP proofs.
33310
44f9665c2091 proper header;
wenzelm
parents: 33243
diff changeset
     7
*)
44f9665c2091 proper header;
wenzelm
parents: 33243
diff changeset
     8
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43062
diff changeset
     9
signature ATP_RECONSTRUCT =
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    10
sig
43678
56d352659500 improved translation of lambdas in THF
nik
parents: 43481
diff changeset
    11
  type ('a, 'b) ho_term = ('a, 'b) ATP_Problem.ho_term
43127
a3f3b7a0e99e export one more function
blanchet
parents: 43102
diff changeset
    12
  type ('a, 'b, 'c) formula = ('a, 'b, 'c) ATP_Problem.formula
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
    13
  type 'a proof = 'a ATP_Proof.proof
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43062
diff changeset
    14
  type locality = ATP_Translate.locality
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    15
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    16
  datatype reconstructor =
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    17
    Metis |
43228
2ed2f092e990 obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents: 43212
diff changeset
    18
    Metis_Full_Types |
2ed2f092e990 obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents: 43212
diff changeset
    19
    Metis_No_Types |
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    20
    SMT of string
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    21
43050
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43039
diff changeset
    22
  datatype play =
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43039
diff changeset
    23
    Played of reconstructor * Time.time |
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
    24
    Trust_Playable of reconstructor * Time.time option |
43166
68e3cd19fee8 show what failed to play
blanchet
parents: 43163
diff changeset
    25
    Failed_to_Play of reconstructor
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    26
36281
dbbf4d5d584d pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents: 36231
diff changeset
    27
  type minimize_command = string list -> string
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    28
  type one_line_params =
43050
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43039
diff changeset
    29
    play * string * (string * locality) list * minimize_command * int * int
38818
61cf050f8b2e improve SPASS hack, when a clause comes from several facts
blanchet
parents: 38752
diff changeset
    30
  type isar_params =
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
    31
    bool * bool * int * string Symtab.table * int list list * int
43102
9a42899ec169 tuned name
blanchet
parents: 43095
diff changeset
    32
    * (string * locality) list vector * int Symtab.table * string proof * thm
42451
a75fcd103cbb automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents: 42449
diff changeset
    33
  val used_facts_in_atp_proof :
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
    34
    Proof.context -> int -> (string * locality) list vector -> string proof
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
    35
    -> (string * locality) list
42876
e336ef6313aa more informative message when Sledgehammer finds an unsound proof
blanchet
parents: 42761
diff changeset
    36
  val used_facts_in_unsound_atp_proof :
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
    37
    Proof.context -> int list list -> int -> (string * locality) list vector
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
    38
    -> 'a proof -> string list option
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    39
  val uses_typed_helpers : int list -> 'a proof -> bool
45378
67ed44d7c929 more detailed preplay output
blanchet
parents: 45209
diff changeset
    40
  val name_of_reconstructor : reconstructor -> string
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    41
  val one_line_proof_text : one_line_params -> string
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
    42
  val make_tvar : string -> typ
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
    43
  val make_tfree : Proof.context -> string -> typ
43094
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
    44
  val term_from_atp :
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
    45
    Proof.context -> bool -> int Symtab.table -> typ option
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
    46
    -> (string, string) ho_term -> term
43127
a3f3b7a0e99e export one more function
blanchet
parents: 43102
diff changeset
    47
  val prop_from_atp :
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
    48
    Proof.context -> bool -> int Symtab.table
43678
56d352659500 improved translation of lambdas in THF
nik
parents: 43481
diff changeset
    49
    -> (string, string, (string, string) ho_term) formula -> term
42968
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
    50
  val isar_proof_text :
43062
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
    51
    Proof.context -> bool -> isar_params -> one_line_params -> string
42968
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
    52
  val proof_text :
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    53
    Proof.context -> bool -> isar_params -> one_line_params -> string
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    54
end;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    55
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43062
diff changeset
    56
structure ATP_Reconstruct : ATP_RECONSTRUCT =
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    57
struct
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    58
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43062
diff changeset
    59
open ATP_Util
38028
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
    60
open ATP_Problem
39452
70a57e40f795 factored out TSTP/SPASS/Vampire proof parsing;
blanchet
parents: 39425
diff changeset
    61
open ATP_Proof
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43062
diff changeset
    62
open ATP_Translate
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    63
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    64
datatype reconstructor =
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    65
  Metis |
43228
2ed2f092e990 obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents: 43212
diff changeset
    66
  Metis_Full_Types |
2ed2f092e990 obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents: 43212
diff changeset
    67
  Metis_No_Types |
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    68
  SMT of string
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    69
43050
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43039
diff changeset
    70
datatype play =
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43039
diff changeset
    71
  Played of reconstructor * Time.time |
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    72
  Trust_Playable of reconstructor * Time.time option |
43166
68e3cd19fee8 show what failed to play
blanchet
parents: 43163
diff changeset
    73
  Failed_to_Play of reconstructor
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    74
36281
dbbf4d5d584d pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents: 36231
diff changeset
    75
type minimize_command = string list -> string
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
    76
type one_line_params =
43050
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43039
diff changeset
    77
  play * string * (string * locality) list * minimize_command * int * int
38818
61cf050f8b2e improve SPASS hack, when a clause comes from several facts
blanchet
parents: 38752
diff changeset
    78
type isar_params =
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
    79
  bool * bool * int * string Symtab.table * int list list * int
43037
ade5c84f860f cleanup proof text generation code
blanchet
parents: 43034
diff changeset
    80
  * (string * locality) list vector * int Symtab.table * string proof * thm
36281
dbbf4d5d584d pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents: 36231
diff changeset
    81
42881
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
    82
val is_typed_helper_name =
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
    83
  String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
    84
39500
d91ef7fbc500 move functions around
blanchet
parents: 39495
diff changeset
    85
fun find_first_in_list_vector vec key =
d91ef7fbc500 move functions around
blanchet
parents: 39495
diff changeset
    86
  Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
d91ef7fbc500 move functions around
blanchet
parents: 39495
diff changeset
    87
                 | (_, value) => value) NONE vec
d91ef7fbc500 move functions around
blanchet
parents: 39495
diff changeset
    88
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
    89
val unprefix_fact_number = space_implode "_" o tl o space_explode "_"
42180
a6c141925a8a added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents: 41742
diff changeset
    90
a6c141925a8a added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents: 41742
diff changeset
    91
val vampire_step_prefix = "f" (* grrr... *)
41203
1393514094d7 fixed more issues with the Vampire output parser, and added support for Vampire's TSTP output (--proof tptp)
blanchet
parents: 41151
diff changeset
    92
42881
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
    93
val extract_step_number =
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
    94
  Int.fromString o perhaps (try (unprefix vampire_step_prefix))
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
    95
43481
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
    96
fun resolve_one_named_fact fact_names s =
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
    97
  case try (unprefix fact_prefix) s of
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
    98
    SOME s' =>
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
    99
    let val s' = s' |> unprefix_fact_number |> unascii_of in
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
   100
      s' |> find_first_in_list_vector fact_names |> Option.map (pair s')
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
   101
    end
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
   102
  | NONE => NONE
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
   103
fun resolve_fact _ fact_names (_, SOME ss) =
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
   104
    map_filter (resolve_one_named_fact fact_names) ss
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   105
  | resolve_fact facts_offset fact_names (num, NONE) =
42881
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   106
    (case extract_step_number num of
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   107
       SOME j =>
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   108
       let val j = j - facts_offset in
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   109
         if j > 0 andalso j <= Vector.length fact_names then
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   110
           Vector.sub (fact_names, j - 1)
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   111
         else
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   112
           []
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   113
       end
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   114
     | NONE => [])
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   115
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   116
fun is_fact conjecture_shape = not o null o resolve_fact 0 conjecture_shape
39453
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   117
43481
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
   118
fun resolve_one_named_conjecture s =
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
   119
  case try (unprefix conjecture_prefix) s of
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
   120
    SOME s' => Int.fromString s'
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
   121
  | NONE => NONE
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
   122
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
   123
fun resolve_conjecture _ (_, SOME ss) =
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
   124
    map_filter resolve_one_named_conjecture ss
42751
f42fd9754724 fixed conjecture resolution bug for Vampire 1.0's TSTP output
blanchet
parents: 42680
diff changeset
   125
  | resolve_conjecture conjecture_shape (num, NONE) =
42881
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   126
    case extract_step_number num of
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   127
      SOME i => (case find_index (exists (curry (op =) i)) conjecture_shape of
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   128
                   ~1 => []
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   129
                 | j => [j])
42751
f42fd9754724 fixed conjecture resolution bug for Vampire 1.0's TSTP output
blanchet
parents: 42680
diff changeset
   130
    | NONE => []
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   131
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   132
fun is_conjecture conjecture_shape =
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   133
  not o null o resolve_conjecture conjecture_shape
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   134
43481
51857e7fa64b clean up SPASS FLOTTER hack
blanchet
parents: 43421
diff changeset
   135
fun is_typed_helper _ (_, SOME ss) = exists is_typed_helper_name ss
42881
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   136
  | is_typed_helper typed_helpers (num, NONE) =
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   137
    (case extract_step_number num of
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   138
       SOME i => member (op =) typed_helpers i
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   139
     | NONE => false)
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   140
42968
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
   141
val leo2_ext = "extcnf_equal_neg"
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
   142
val isa_ext = Thm.get_name_hint @{thm ext}
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
   143
val isa_short_ext = Long_Name.base_name isa_ext
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
   144
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
   145
fun ext_name ctxt =
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
   146
  if Thm.eq_thm_prop (@{thm ext},
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
   147
         singleton (Attrib.eval_thms ctxt) (Facts.named isa_short_ext, [])) then
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
   148
    isa_short_ext
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
   149
  else
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
   150
    isa_ext
39453
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   151
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   152
fun add_fact _ facts_offset fact_names (Inference (name, _, _, [])) =
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   153
    union (op =) (resolve_fact facts_offset fact_names name)
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   154
  | add_fact ctxt _ _ (Inference (_, _, rule, _)) =
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   155
    if rule = leo2_ext then insert (op =) (ext_name ctxt, General) else I
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   156
  | add_fact _ _ _ _ = I
42968
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
   157
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   158
fun used_facts_in_atp_proof ctxt facts_offset fact_names atp_proof =
43296
755e3d5ea3f2 avoid duplicate facts, which confuse the minimizer output
blanchet
parents: 43291
diff changeset
   159
  if null atp_proof then Vector.foldl (uncurry (union (op =))) [] fact_names
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   160
  else fold (add_fact ctxt facts_offset fact_names) atp_proof []
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   161
43291
9f33b4ec866c don't generate unsound proof error for missing proofs
blanchet
parents: 43233
diff changeset
   162
fun is_conjecture_used_in_proof conjecture_shape =
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   163
  exists (fn Inference (name, _, _, []) => is_conjecture conjecture_shape name
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   164
           | _ => false)
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   165
44399
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44396
diff changeset
   166
(* (quasi-)underapproximation of the truth *)
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44396
diff changeset
   167
fun is_locality_global Local = false
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44396
diff changeset
   168
  | is_locality_global Assum = false
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44396
diff changeset
   169
  | is_locality_global Chained = false
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44396
diff changeset
   170
  | is_locality_global _ = true
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44396
diff changeset
   171
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   172
fun used_facts_in_unsound_atp_proof _ _ _ _ [] = NONE
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   173
  | used_facts_in_unsound_atp_proof ctxt conjecture_shape facts_offset
42876
e336ef6313aa more informative message when Sledgehammer finds an unsound proof
blanchet
parents: 42761
diff changeset
   174
                                    fact_names atp_proof =
43291
9f33b4ec866c don't generate unsound proof error for missing proofs
blanchet
parents: 43233
diff changeset
   175
    let
9f33b4ec866c don't generate unsound proof error for missing proofs
blanchet
parents: 43233
diff changeset
   176
      val used_facts =
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   177
        used_facts_in_atp_proof ctxt facts_offset fact_names atp_proof
43291
9f33b4ec866c don't generate unsound proof error for missing proofs
blanchet
parents: 43233
diff changeset
   178
    in
9f33b4ec866c don't generate unsound proof error for missing proofs
blanchet
parents: 43233
diff changeset
   179
      if forall (is_locality_global o snd) used_facts andalso
9f33b4ec866c don't generate unsound proof error for missing proofs
blanchet
parents: 43233
diff changeset
   180
         not (is_conjecture_used_in_proof conjecture_shape atp_proof) then
9f33b4ec866c don't generate unsound proof error for missing proofs
blanchet
parents: 43233
diff changeset
   181
        SOME (map fst used_facts)
9f33b4ec866c don't generate unsound proof error for missing proofs
blanchet
parents: 43233
diff changeset
   182
      else
9f33b4ec866c don't generate unsound proof error for missing proofs
blanchet
parents: 43233
diff changeset
   183
        NONE
9f33b4ec866c don't generate unsound proof error for missing proofs
blanchet
parents: 43233
diff changeset
   184
    end
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   185
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   186
fun uses_typed_helpers typed_helpers =
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   187
  exists (fn Inference (name, _, _, []) => is_typed_helper typed_helpers name
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   188
           | _ => false)
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   189
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   190
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   191
(** Soft-core proof reconstruction: Metis one-liner **)
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   192
43228
2ed2f092e990 obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents: 43212
diff changeset
   193
(* unfortunate leaking abstraction *)
43233
2749c357f865 use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
blanchet
parents: 43228
diff changeset
   194
fun name_of_reconstructor Metis = "metis"
2749c357f865 use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
blanchet
parents: 43228
diff changeset
   195
  | name_of_reconstructor Metis_Full_Types = "metis (full_types)"
2749c357f865 use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
blanchet
parents: 43228
diff changeset
   196
  | name_of_reconstructor Metis_No_Types = "metis (no_types)"
2749c357f865 use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
blanchet
parents: 43228
diff changeset
   197
  | name_of_reconstructor (SMT _) = "smt"
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   198
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   199
fun reconstructor_settings (SMT settings) = settings
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   200
  | reconstructor_settings _ = ""
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   201
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   202
fun string_for_label (s, num) = s ^ string_of_int num
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   203
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   204
fun show_time NONE = ""
43034
18259246abb5 try both "metis" and (on failure) "metisFT" in replay
blanchet
parents: 43033
diff changeset
   205
  | show_time (SOME ext_time) = " (" ^ string_from_ext_time ext_time ^ ")"
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   206
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   207
fun set_settings "" = ""
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   208
  | set_settings settings = "using [[" ^ settings ^ "]] "
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   209
fun apply_on_subgoal settings _ 1 = set_settings settings ^ "by "
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   210
  | apply_on_subgoal settings 1 _ = set_settings settings ^ "apply "
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   211
  | apply_on_subgoal settings i n =
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   212
    "prefer " ^ string_of_int i ^ " " ^ apply_on_subgoal settings 1 n
44002
ae53f1304ad5 put parentheses around non-trivial metis call
blanchet
parents: 43907
diff changeset
   213
fun command_call name [] =
ae53f1304ad5 put parentheses around non-trivial metis call
blanchet
parents: 43907
diff changeset
   214
    name |> not (Lexicon.is_identifier name) ? enclose "(" ")"
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   215
  | command_call name args = "(" ^ name ^ " " ^ space_implode " " args ^ ")"
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   216
fun try_command_line banner time command =
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   217
  banner ^ ": " ^ Markup.markup Markup.sendback command ^ show_time time ^ "."
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   218
fun using_labels [] = ""
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   219
  | using_labels ls =
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   220
    "using " ^ space_implode " " (map string_for_label ls) ^ " "
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   221
fun reconstructor_command reconstructor i n (ls, ss) =
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   222
  using_labels ls ^
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   223
  apply_on_subgoal (reconstructor_settings reconstructor) i n ^
43233
2749c357f865 use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
blanchet
parents: 43228
diff changeset
   224
  command_call (name_of_reconstructor reconstructor) ss
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   225
fun minimize_line _ [] = ""
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   226
  | minimize_line minimize_command ss =
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   227
    case minimize_command ss of
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   228
      "" => ""
43006
ff631c45797e make output more concise
blanchet
parents: 43004
diff changeset
   229
    | command => "\nTo minimize: " ^ Markup.markup Markup.sendback command ^ "."
39453
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   230
40723
a82badd0e6ef put facts found by SMT solver in alphabetical order and omit chained facts, as was done already for ATP proofs
blanchet
parents: 40627
diff changeset
   231
val split_used_facts =
a82badd0e6ef put facts found by SMT solver in alphabetical order and omit chained facts, as was done already for ATP proofs
blanchet
parents: 40627
diff changeset
   232
  List.partition (curry (op =) Chained o snd)
39453
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   233
  #> pairself (sort_distinct (string_ord o pairself fst))
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   234
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   235
fun one_line_proof_text (preplay, banner, used_facts, minimize_command,
43037
ade5c84f860f cleanup proof text generation code
blanchet
parents: 43034
diff changeset
   236
                         subgoal, subgoal_count) =
39453
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   237
  let
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   238
    val (chained, extra) = split_used_facts used_facts
43166
68e3cd19fee8 show what failed to play
blanchet
parents: 43163
diff changeset
   239
    val (failed, reconstructor, ext_time) =
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   240
      case preplay of
43050
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43039
diff changeset
   241
        Played (reconstructor, time) =>
43166
68e3cd19fee8 show what failed to play
blanchet
parents: 43163
diff changeset
   242
        (false, reconstructor, (SOME (false, time)))
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   243
      | Trust_Playable (reconstructor, time) =>
43166
68e3cd19fee8 show what failed to play
blanchet
parents: 43163
diff changeset
   244
        (false, reconstructor,
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   245
         case time of
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   246
           NONE => NONE
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   247
         | SOME time =>
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   248
           if time = Time.zeroTime then NONE else SOME (true, time))
43166
68e3cd19fee8 show what failed to play
blanchet
parents: 43163
diff changeset
   249
      | Failed_to_Play reconstructor => (true, reconstructor, NONE)
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   250
    val try_line =
43166
68e3cd19fee8 show what failed to play
blanchet
parents: 43163
diff changeset
   251
      ([], map fst extra)
68e3cd19fee8 show what failed to play
blanchet
parents: 43163
diff changeset
   252
      |> reconstructor_command reconstructor subgoal subgoal_count
68e3cd19fee8 show what failed to play
blanchet
parents: 43163
diff changeset
   253
      |> (if failed then enclose "One-line proof reconstruction failed: " "."
68e3cd19fee8 show what failed to play
blanchet
parents: 43163
diff changeset
   254
          else try_command_line banner ext_time)
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   255
  in try_line ^ minimize_line minimize_command (map fst (extra @ chained)) end
39453
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   256
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   257
(** Hard-core proof reconstruction: structured Isar proofs **)
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   258
39425
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   259
fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   260
fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda v t
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   261
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
   262
fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS)
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   263
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
   264
  let val ww = "'" ^ w in
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   265
    TFree (ww, the_default HOLogic.typeS (Variable.def_sort ctxt (ww, ~1)))
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   266
  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
   267
39453
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   268
val indent_size = 2
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   269
val no_label = ("", ~1)
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   270
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   271
val raw_prefix = "X"
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   272
val assum_prefix = "A"
42180
a6c141925a8a added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents: 41742
diff changeset
   273
val have_prefix = "F"
39370
f8292d3020db use same hack as in "Async_Manager" to work around Proof General bug
blanchet
parents: 39368
diff changeset
   274
39453
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   275
fun raw_label_for_name conjecture_shape name =
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   276
  case resolve_conjecture conjecture_shape name of
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   277
    [j] => (conjecture_prefix, j)
39455
c6b21584f336 merge constructors
blanchet
parents: 39454
diff changeset
   278
  | _ => case Int.fromString (fst name) of
39453
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   279
           SOME j => (raw_prefix, j)
39455
c6b21584f336 merge constructors
blanchet
parents: 39454
diff changeset
   280
         | NONE => (raw_prefix ^ fst name, 0)
39453
1740a2d6bef9 use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
blanchet
parents: 39452
diff changeset
   281
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   282
(**** INTERPRETATION OF TSTP SYNTAX TREES ****)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   283
43678
56d352659500 improved translation of lambdas in THF
nik
parents: 43481
diff changeset
   284
exception HO_TERM of (string, string) ho_term list
56d352659500 improved translation of lambdas in THF
nik
parents: 43481
diff changeset
   285
exception FORMULA of (string, string, (string, string) ho_term) formula list
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   286
exception SAME of unit
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   287
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   288
(* Type variables are given the basic sort "HOL.type". Some will later be
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   289
   constrained by information from type literals, or by type inference. *)
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
   290
fun typ_from_atp ctxt (u as ATerm (a, us)) =
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   291
  let val Ts = map (typ_from_atp ctxt) us in
38748
69fea359d3f8 renaming
blanchet
parents: 38738
diff changeset
   292
    case strip_prefix_and_unascii type_const_prefix a of
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   293
      SOME b => Type (invert_const b, Ts)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   294
    | NONE =>
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   295
      if not (null us) then
43678
56d352659500 improved translation of lambdas in THF
nik
parents: 43481
diff changeset
   296
        raise HO_TERM [u]  (* only "tconst"s have type arguments *)
38748
69fea359d3f8 renaming
blanchet
parents: 38738
diff changeset
   297
      else case strip_prefix_and_unascii tfree_prefix a of
43135
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   298
        SOME b => make_tfree ctxt b
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   299
      | NONE =>
43302
566f970006e5 improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents: 43296
diff changeset
   300
        (* Could be an Isabelle variable or a variable from the ATP, say "X1"
566f970006e5 improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents: 43296
diff changeset
   301
           or "_5018". Sometimes variables from the ATP are indistinguishable
566f970006e5 improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents: 43296
diff changeset
   302
           from Isabelle variables, which forces us to use a type parameter in
566f970006e5 improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents: 43296
diff changeset
   303
           all cases. *)
566f970006e5 improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents: 43296
diff changeset
   304
        (a |> perhaps (strip_prefix_and_unascii tvar_prefix), HOLogic.typeS)
566f970006e5 improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
blanchet
parents: 43296
diff changeset
   305
        |> Type_Infer.param 0
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   306
  end
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   307
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   308
(* Type class literal applied to a type. Returns triple of polarity, class,
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   309
   type. *)
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
   310
fun type_constraint_from_term ctxt (u as ATerm (a, us)) =
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   311
  case (strip_prefix_and_unascii class_prefix a, map (typ_from_atp ctxt) us) of
42606
0c76cf483899 show sorts not just types in Isar proofs + tuning
blanchet
parents: 42595
diff changeset
   312
    (SOME b, [T]) => (b, T)
43678
56d352659500 improved translation of lambdas in THF
nik
parents: 43481
diff changeset
   313
  | _ => raise HO_TERM [u]
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   314
43907
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   315
(* Accumulate type constraints in a formula: negative type literals. *)
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   316
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
   317
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
   318
  | 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
   319
  | add_type_constraint _ _ = I
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   320
43094
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
   321
fun repair_variable_name f s =
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   322
  let
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   323
    fun subscript_name s n = s ^ nat_subscript n
38488
3abda3c76df9 handle E's Skolem constants more gracefully
blanchet
parents: 38282
diff changeset
   324
    val s = String.map f s
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   325
  in
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   326
    case space_explode "_" s of
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   327
      [_] => (case take_suffix Char.isDigit (String.explode s) of
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   328
                (cs1 as _ :: _, cs2 as _ :: _) =>
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   329
                subscript_name (String.implode cs1)
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   330
                               (the (Int.fromString (String.implode cs2)))
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   331
              | (_, _) => s)
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   332
    | [s1, s2] => (case Int.fromString s2 of
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   333
                     SOME n => subscript_name s1 n
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   334
                   | NONE => s)
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   335
    | _ => s
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   336
  end
43182
649bada59658 slacker version of "fastype_of", in case a function has dummy type
blanchet
parents: 43176
diff changeset
   337
43907
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   338
(* The number of type arguments of a constant, zero if it's monomorphic. For
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   339
   (instances of) Skolem pseudoconstants, this information is encoded in the
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   340
   constant name. *)
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   341
fun num_type_args thy s =
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   342
  if String.isPrefix skolem_const_prefix s then
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   343
    s |> space_explode Long_Name.separator |> List.last |> Int.fromString |> the
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   344
  else
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   345
    (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
   346
43182
649bada59658 slacker version of "fastype_of", in case a function has dummy type
blanchet
parents: 43176
diff changeset
   347
fun slack_fastype_of t = fastype_of t handle TERM _ => HOLogic.typeT
649bada59658 slacker version of "fastype_of", in case a function has dummy type
blanchet
parents: 43176
diff changeset
   348
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   349
(* First-order translation. No types are known for variables. "HOLogic.typeT"
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   350
   should allow them to be inferred. *)
43135
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   351
fun term_from_atp ctxt textual sym_tab =
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   352
  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
   353
    val thy = Proof_Context.theory_of ctxt
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43204
diff changeset
   354
    (* For Metis, we use 1 rather than 0 because variable references in clauses
050a03afe024 Metis code cleanup
blanchet
parents: 43204
diff changeset
   355
       may otherwise conflict with variable constraints in the goal. At least,
050a03afe024 Metis code cleanup
blanchet
parents: 43204
diff changeset
   356
       type inference often fails otherwise. See also "axiom_inference" in
050a03afe024 Metis code cleanup
blanchet
parents: 43204
diff changeset
   357
       "Metis_Reconstruct". *)
43094
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
   358
    val var_index = if textual then 0 else 1
43131
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   359
    fun do_term extra_ts opt_T u =
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   360
      case u of
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   361
        ATerm (s, us) =>
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   362
        if String.isPrefix simple_type_prefix s then
42962
3b50fdeb6cfc started adding support for THF output (but no lambdas)
blanchet
parents: 42943
diff changeset
   363
          @{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
   364
        else if s = tptp_equal then
43093
blanchet
parents: 43085
diff changeset
   365
          let val ts = map (do_term [] NONE) us in
43094
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
   366
            if textual andalso length ts = 2 andalso
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
   367
              hd ts aconv List.last ts then
39106
5ab6a3707499 fix trivial "x = x" fact detection
blanchet
parents: 38988
diff changeset
   368
              (* Vampire is keen on producing these. *)
5ab6a3707499 fix trivial "x = x" fact detection
blanchet
parents: 38988
diff changeset
   369
              @{const True}
5ab6a3707499 fix trivial "x = x" fact detection
blanchet
parents: 38988
diff changeset
   370
            else
5ab6a3707499 fix trivial "x = x" fact detection
blanchet
parents: 38988
diff changeset
   371
              list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT), ts)
5ab6a3707499 fix trivial "x = x" fact detection
blanchet
parents: 38988
diff changeset
   372
          end
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   373
        else case strip_prefix_and_unascii const_prefix s of
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   374
          SOME s' =>
42761
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   375
          let
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   376
            val ((s', s''), mangled_us) =
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   377
              s' |> unmangled_const |>> `invert_const
42761
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   378
          in
42755
4603154a3018 robustly detect how many type args were passed to the ATP, even if some of them were omitted
blanchet
parents: 42751
diff changeset
   379
            if s' = type_tag_name then
42589
9f7c48463645 restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents: 42587
diff changeset
   380
              case mangled_us @ us of
9f7c48463645 restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents: 42587
diff changeset
   381
                [typ_u, term_u] =>
43135
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   382
                do_term extra_ts (SOME (typ_from_atp ctxt typ_u)) term_u
43678
56d352659500 improved translation of lambdas in THF
nik
parents: 43481
diff changeset
   383
              | _ => raise HO_TERM us
42966
4e2d6c1e5392 more work on parsing LEO-II proofs without lambdas
blanchet
parents: 42962
diff changeset
   384
            else if s' = predicator_name then
43093
blanchet
parents: 43085
diff changeset
   385
              do_term [] (SOME @{typ bool}) (hd us)
42966
4e2d6c1e5392 more work on parsing LEO-II proofs without lambdas
blanchet
parents: 42962
diff changeset
   386
            else if s' = app_op_name then
43131
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   387
              let val extra_t = do_term [] NONE (List.last us) in
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   388
                do_term (extra_t :: extra_ts)
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   389
                        (case opt_T of
43182
649bada59658 slacker version of "fastype_of", in case a function has dummy type
blanchet
parents: 43176
diff changeset
   390
                           SOME T => SOME (slack_fastype_of extra_t --> T)
43131
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   391
                         | NONE => NONE)
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   392
                        (nth us (length us - 2))
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   393
              end
44396
66b9b3fcd608 started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents: 44121
diff changeset
   394
            else if s' = type_guard_name then
42551
cd99d6d3027a reconstruct TFF type predicates correctly for ToFoF
blanchet
parents: 42549
diff changeset
   395
              @{const True} (* ignore type predicates *)
42549
b9754f26c7bc handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents: 42544
diff changeset
   396
            else
b9754f26c7bc handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents: 42544
diff changeset
   397
              let
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   398
                val new_skolem = String.isPrefix new_skolem_const_prefix s''
42755
4603154a3018 robustly detect how many type args were passed to the ATP, even if some of them were omitted
blanchet
parents: 42751
diff changeset
   399
                val num_ty_args =
4603154a3018 robustly detect how many type args were passed to the ATP, even if some of them were omitted
blanchet
parents: 42751
diff changeset
   400
                  length us - the_default 0 (Symtab.lookup sym_tab s)
42549
b9754f26c7bc handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents: 42544
diff changeset
   401
                val (type_us, term_us) =
b9754f26c7bc handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents: 42544
diff changeset
   402
                  chop num_ty_args us |>> append mangled_us
43093
blanchet
parents: 43085
diff changeset
   403
                val term_ts = map (do_term [] NONE) term_us
42549
b9754f26c7bc handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents: 42544
diff changeset
   404
                val T =
43183
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   405
                  (if not (null type_us) andalso
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   406
                      num_type_args thy s' = length type_us then
43191
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   407
                     let val Ts = type_us |> map (typ_from_atp ctxt) in
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   408
                       if new_skolem then
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   409
                         SOME (Type_Infer.paramify_vars (tl Ts ---> hd Ts))
43200
ca7b0a48515d imported patch metis_reconstr_give_type_infer_a_chance
blanchet
parents: 43199
diff changeset
   410
                       else if textual then
ca7b0a48515d imported patch metis_reconstr_give_type_infer_a_chance
blanchet
parents: 43199
diff changeset
   411
                         try (Sign.const_instance thy) (s', Ts)
43191
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   412
                       else
43200
ca7b0a48515d imported patch metis_reconstr_give_type_infer_a_chance
blanchet
parents: 43199
diff changeset
   413
                         NONE
43191
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   414
                     end
43183
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   415
                   else
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   416
                     NONE)
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   417
                  |> (fn SOME T => T
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   418
                       | NONE => map slack_fastype_of term_ts --->
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   419
                                 (case opt_T of
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   420
                                    SOME T => T
faece9668bce don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
blanchet
parents: 43182
diff changeset
   421
                                  | NONE => HOLogic.typeT))
43191
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   422
                val t =
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   423
                  if new_skolem then
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   424
                    Var ((new_skolem_var_name_from_const s'', var_index), T)
43191
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   425
                  else
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   426
                    Const (unproxify_const s', T)
0a72c0527111 fixed reconstruction of new Skolem constants in new Metis
blanchet
parents: 43184
diff changeset
   427
              in list_comb (t, term_ts @ extra_ts) end
42549
b9754f26c7bc handle special constants correctly in Isar proof reconstruction code, especially type predicates
blanchet
parents: 42544
diff changeset
   428
          end
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   429
        | NONE => (* a free or schematic variable *)
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   430
          let
45042
89341b897412 better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents: 44783
diff changeset
   431
            val term_ts = map (do_term [] NONE) us
89341b897412 better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents: 44783
diff changeset
   432
            val ts = term_ts @ extra_ts
89341b897412 better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents: 44783
diff changeset
   433
            val T =
89341b897412 better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents: 44783
diff changeset
   434
              case opt_T of
89341b897412 better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents: 44783
diff changeset
   435
                SOME T => map slack_fastype_of term_ts ---> T
89341b897412 better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents: 44783
diff changeset
   436
              | NONE => map slack_fastype_of ts ---> HOLogic.typeT
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   437
            val t =
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   438
              case strip_prefix_and_unascii fixed_var_prefix s of
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   439
                SOME s =>
43907
073ab5379842 pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents: 43863
diff changeset
   440
                (* FIXME: reconstruction of lambda-lifting *)
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   441
                Free (s, T)
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   442
              | NONE =>
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   443
                case strip_prefix_and_unascii schematic_var_prefix s of
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   444
                  SOME s => Var ((s, var_index), T)
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   445
                | NONE =>
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   446
                  Var ((s |> textual ? repair_variable_name Char.toLower,
43095
ccf1c09dea82 more robust and simpler adjustment computation
blanchet
parents: 43094
diff changeset
   447
                        var_index), T)
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   448
          in list_comb (t, ts) end
43093
blanchet
parents: 43085
diff changeset
   449
  in do_term [] end
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   450
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
   451
fun term_from_atom ctxt textual sym_tab pos (u as ATerm (s, _)) =
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   452
  if String.isPrefix class_prefix s then
43135
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   453
    add_type_constraint pos (type_constraint_from_term ctxt u)
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   454
    #> pair @{const True}
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   455
  else
43135
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   456
    pair (term_from_atp ctxt textual sym_tab (SOME @{typ bool}) u)
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   457
36555
8ff45c2076da expand combinators in Isar proofs constructed by Sledgehammer;
blanchet
parents: 36551
diff changeset
   458
val combinator_table =
39953
aa54f347e5e2 hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
blanchet
parents: 39710
diff changeset
   459
  [(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def_raw}),
aa54f347e5e2 hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
blanchet
parents: 39710
diff changeset
   460
   (@{const_name Meson.COMBK}, @{thm Meson.COMBK_def_raw}),
aa54f347e5e2 hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
blanchet
parents: 39710
diff changeset
   461
   (@{const_name Meson.COMBB}, @{thm Meson.COMBB_def_raw}),
aa54f347e5e2 hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
blanchet
parents: 39710
diff changeset
   462
   (@{const_name Meson.COMBC}, @{thm Meson.COMBC_def_raw}),
aa54f347e5e2 hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
blanchet
parents: 39710
diff changeset
   463
   (@{const_name Meson.COMBS}, @{thm Meson.COMBS_def_raw})]
36555
8ff45c2076da expand combinators in Isar proofs constructed by Sledgehammer;
blanchet
parents: 36551
diff changeset
   464
42761
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   465
fun uncombine_term thy =
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   466
  let
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   467
    fun aux (t1 $ t2) = betapply (pairself aux (t1, t2))
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   468
      | aux (Abs (s, T, t')) = Abs (s, T, aux t')
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   469
      | aux (t as Const (x as (s, _))) =
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   470
        (case AList.lookup (op =) combinator_table s of
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   471
           SOME thm => thm |> prop_of |> specialize_type thy x
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   472
                           |> Logic.dest_equals |> snd
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   473
         | NONE => t)
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   474
      | aux t = t
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   475
  in aux end
36555
8ff45c2076da expand combinators in Isar proofs constructed by Sledgehammer;
blanchet
parents: 36551
diff changeset
   476
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   477
(* Update schematic type variables with detected sort constraints. It's not
42563
e70ffe3846d0 improve helper type instantiation code
blanchet
parents: 42562
diff changeset
   478
   totally clear whether this code is necessary. *)
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   479
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
   480
  let
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   481
    fun do_type (Type (a, Ts)) = Type (a, map do_type Ts)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   482
      | do_type (TVar (xi, s)) =
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   483
        TVar (xi, the_default s (Vartab.lookup tvar_tab xi))
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   484
      | do_type (TFree z) = TFree z
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   485
    fun do_term (Const (a, T)) = Const (a, do_type T)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   486
      | do_term (Free (a, T)) = Free (a, do_type T)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   487
      | do_term (Var (xi, T)) = Var (xi, do_type T)
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   488
      | do_term (t as Bound _) = t
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   489
      | 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
   490
      | do_term (t1 $ t2) = do_term t1 $ do_term t2
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   491
  in t |> not (Vartab.is_empty tvar_tab) ? do_term end
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   492
39425
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   493
fun quantify_over_var quant_of var_s t =
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   494
  let
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   495
    val vars = [] |> Term.add_vars t |> filter (fn ((s, _), _) => s = var_s)
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   496
                  |> map Var
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   497
  in fold_rev quant_of vars t end
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   498
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38066
diff changeset
   499
(* Interpret an ATP formula as a HOL term, extracting sort constraints as they
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38066
diff changeset
   500
   appear in the formula. *)
43184
b16693484c5d reveal Skolems in new Metis
blanchet
parents: 43183
diff changeset
   501
fun prop_from_atp ctxt textual sym_tab phi =
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   502
  let
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   503
    fun do_formula pos phi =
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   504
      case phi of
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   505
        AQuant (_, [], phi) => do_formula pos phi
42526
46d485f8d144 added room for types in ATP quantifiers
blanchet
parents: 42451
diff changeset
   506
      | AQuant (q, (s, _) :: xs, phi') =>
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   507
        do_formula pos (AQuant (q, xs, phi'))
42526
46d485f8d144 added room for types in ATP quantifiers
blanchet
parents: 42451
diff changeset
   508
        (* FIXME: TFF *)
39425
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   509
        #>> quantify_over_var (case q of
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   510
                                 AForall => forall_of
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   511
                               | AExists => exists_of)
43095
ccf1c09dea82 more robust and simpler adjustment computation
blanchet
parents: 43094
diff changeset
   512
                              (s |> textual ? repair_variable_name Char.toLower)
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   513
      | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   514
      | AConn (c, [phi1, phi2]) =>
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   515
        do_formula (pos |> c = AImplies ? not) phi1
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   516
        ##>> do_formula pos phi2
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   517
        #>> (case c of
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   518
               AAnd => s_conj
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   519
             | AOr => s_disj
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   520
             | AImplies => s_imp
38038
584ab1a3a523 more robust proof reconstruction
blanchet
parents: 38036
diff changeset
   521
             | AIff => s_iff
43163
31babd4b1552 killed odd connectives
blanchet
parents: 43136
diff changeset
   522
             | ANot => raise Fail "impossible connective")
43135
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   523
      | AAtom tm => term_from_atom ctxt textual sym_tab pos tm
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   524
      | _ => raise FORMULA [phi]
38014
81c23d286f0c extract sort constraints from FOFs properly;
blanchet
parents: 38007
diff changeset
   525
  in repair_tvar_sorts (do_formula true phi Vartab.empty) end
37991
3093ab32f1e7 proof reconstruction for full FOF terms
blanchet
parents: 37962
diff changeset
   526
43131
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   527
fun infer_formula_types ctxt =
39288
f1ae2493d93f eliminated aliases of Type.constraint;
wenzelm
parents: 39134
diff changeset
   528
  Type.constraint HOLogic.boolT
42761
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   529
  #> Syntax.check_term
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   530
         (Proof_Context.set_mode Proof_Context.mode_schematic ctxt)
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   531
43184
b16693484c5d reveal Skolems in new Metis
blanchet
parents: 43183
diff changeset
   532
fun uncombined_etc_prop_from_atp ctxt textual sym_tab =
43136
cf5cda219058 handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
blanchet
parents: 43135
diff changeset
   533
  let val thy = Proof_Context.theory_of ctxt in
43184
b16693484c5d reveal Skolems in new Metis
blanchet
parents: 43183
diff changeset
   534
    prop_from_atp ctxt textual sym_tab
43176
29a3a1a7794d only uncombine combinators in textual Isar proofs, not in Metis
blanchet
parents: 43168
diff changeset
   535
    #> textual ? uncombine_term thy #> infer_formula_types ctxt
43136
cf5cda219058 handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
blanchet
parents: 43135
diff changeset
   536
  end
cf5cda219058 handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
blanchet
parents: 43135
diff changeset
   537
43093
blanchet
parents: 43085
diff changeset
   538
(**** Translation of TSTP files to Isar proofs ****)
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   539
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   540
fun unvarify_term (Var ((s, 0), T)) = Free (s, T)
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   541
  | unvarify_term t = raise TERM ("unvarify_term: non-Var", [t])
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   542
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
   543
fun decode_line sym_tab (Definition (name, phi1, phi2)) ctxt =
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   544
    let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42227
diff changeset
   545
      val thy = Proof_Context.theory_of ctxt
43184
b16693484c5d reveal Skolems in new Metis
blanchet
parents: 43183
diff changeset
   546
      val t1 = prop_from_atp ctxt true sym_tab phi1
36551
cc42df660808 improve unskolemization
blanchet
parents: 36548
diff changeset
   547
      val vars = snd (strip_comb t1)
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   548
      val frees = map unvarify_term vars
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   549
      val unvarify_args = subst_atomic (vars ~~ frees)
43184
b16693484c5d reveal Skolems in new Metis
blanchet
parents: 43183
diff changeset
   550
      val t2 = prop_from_atp ctxt true sym_tab phi2
36551
cc42df660808 improve unskolemization
blanchet
parents: 36548
diff changeset
   551
      val (t1, t2) =
cc42df660808 improve unskolemization
blanchet
parents: 36548
diff changeset
   552
        HOLogic.eq_const HOLogic.typeT $ t1 $ t2
43131
9e9420122f91 fixed interaction between type tags and hAPP in reconstruction code
blanchet
parents: 43130
diff changeset
   553
        |> unvarify_args |> uncombine_term thy |> infer_formula_types ctxt
36555
8ff45c2076da expand combinators in Isar proofs constructed by Sledgehammer;
blanchet
parents: 36551
diff changeset
   554
        |> HOLogic.dest_eq
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   555
    in
39368
f661064b2b80 generalize proof reconstruction code;
blanchet
parents: 39353
diff changeset
   556
      (Definition (name, t1, t2),
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 44002
diff changeset
   557
       fold Variable.declare_term (maps Misc_Legacy.term_frees [t1, t2]) ctxt)
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   558
    end
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   559
  | decode_line sym_tab (Inference (name, u, rule, deps)) ctxt =
43184
b16693484c5d reveal Skolems in new Metis
blanchet
parents: 43183
diff changeset
   560
    let val t = u |> uncombined_etc_prop_from_atp ctxt true sym_tab in
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   561
      (Inference (name, t, rule, deps),
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 44002
diff changeset
   562
       fold Variable.declare_term (Misc_Legacy.term_frees t) ctxt)
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   563
    end
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
   564
fun decode_lines ctxt sym_tab lines =
8c32a0160b0d more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
blanchet
parents: 43131
diff changeset
   565
  fst (fold_map (decode_line sym_tab) lines ctxt)
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   566
38035
0ed953eac020 fix proof reconstruction for latest Vampire
blanchet
parents: 38034
diff changeset
   567
fun is_same_inference _ (Definition _) = false
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   568
  | is_same_inference t (Inference (_, t', _, _)) = t aconv t'
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   569
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   570
(* No "real" literals means only type information (tfree_tcs, clsrel, or
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   571
   clsarity). *)
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   572
val is_only_type_information = curry (op aconv) HOLogic.true_const
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   573
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   574
fun replace_one_dependency (old, new) dep =
42968
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
   575
  if is_same_atp_step dep old then new else [dep]
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   576
fun replace_dependencies_in_line _ (line as Definition _) = line
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   577
  | replace_dependencies_in_line p (Inference (name, t, rule, deps)) =
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   578
    Inference (name, t, rule,
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   579
               fold (union (op =) o replace_one_dependency p) deps [])
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   580
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40114
diff changeset
   581
(* Discard facts; consolidate adjacent lines that prove the same formula, since
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38066
diff changeset
   582
   they differ only in type information.*)
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   583
fun add_line _ _ (line as Definition _) lines = line :: lines
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   584
  | add_line conjecture_shape fact_names (Inference (name, t, rule, [])) lines =
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40114
diff changeset
   585
    (* No dependencies: fact, conjecture, or (for Vampire) internal facts or
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38066
diff changeset
   586
       definitions. *)
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   587
    if is_fact fact_names name then
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40114
diff changeset
   588
      (* Facts are not proof lines. *)
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   589
      if is_only_type_information t then
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   590
        map (replace_dependencies_in_line (name, [])) lines
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   591
      (* Is there a repetition? If so, replace later line by earlier one. *)
38035
0ed953eac020 fix proof reconstruction for latest Vampire
blanchet
parents: 38034
diff changeset
   592
      else case take_prefix (not o is_same_inference t) lines of
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   593
        (_, []) => lines (* no repetition of proof line *)
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   594
      | (pre, Inference (name', _, _, _) :: post) =>
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   595
        pre @ map (replace_dependencies_in_line (name', [name])) post
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40068
diff changeset
   596
      | _ => raise Fail "unexpected inference"
39370
f8292d3020db use same hack as in "Async_Manager" to work around Proof General bug
blanchet
parents: 39368
diff changeset
   597
    else if is_conjecture conjecture_shape name then
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   598
      Inference (name, s_not t, rule, []) :: lines
36551
cc42df660808 improve unskolemization
blanchet
parents: 36548
diff changeset
   599
    else
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   600
      map (replace_dependencies_in_line (name, [])) lines
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   601
  | add_line _ _ (Inference (name, t, rule, deps)) lines =
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   602
    (* Type information will be deleted later; skip repetition test. *)
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   603
    if is_only_type_information t then
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   604
      Inference (name, t, rule, deps) :: lines
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   605
    (* Is there a repetition? If so, replace later line by earlier one. *)
38035
0ed953eac020 fix proof reconstruction for latest Vampire
blanchet
parents: 38034
diff changeset
   606
    else case take_prefix (not o is_same_inference t) lines of
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   607
      (* FIXME: Doesn't this code risk conflating proofs involving different
38035
0ed953eac020 fix proof reconstruction for latest Vampire
blanchet
parents: 38034
diff changeset
   608
         types? *)
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   609
       (_, []) => Inference (name, t, rule, deps) :: lines
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   610
     | (pre, Inference (name', t', rule, _) :: post) =>
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   611
       Inference (name, t', rule, deps) ::
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   612
       pre @ map (replace_dependencies_in_line (name', [name])) post
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40068
diff changeset
   613
     | _ => raise Fail "unexpected inference"
22044
6c0702a96076 More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents: 22012
diff changeset
   614
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   615
(* Recursively delete empty lines (type information) from the proof. *)
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   616
fun add_nontrivial_line (line as Inference (name, t, _, [])) lines =
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   617
    if is_only_type_information t then delete_dependency name lines
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   618
    else line :: lines
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   619
  | add_nontrivial_line line lines = line :: lines
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   620
and delete_dependency name lines =
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   621
  fold_rev add_nontrivial_line
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   622
           (map (replace_dependencies_in_line (name, [])) lines) []
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   623
37323
2f2f0d246d0f handle Vampire's definitions smoothly
blanchet
parents: 37322
diff changeset
   624
(* ATPs sometimes reuse free variable names in the strangest ways. Removing
2f2f0d246d0f handle Vampire's definitions smoothly
blanchet
parents: 37322
diff changeset
   625
   offending lines often does the trick. *)
36560
45c1870f234f fixed definition of "bad frees" so that it actually works
blanchet
parents: 36559
diff changeset
   626
fun is_bad_free frees (Free x) = not (member (op =) frees x)
45c1870f234f fixed definition of "bad frees" so that it actually works
blanchet
parents: 36559
diff changeset
   627
  | is_bad_free _ _ = false
22470
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   628
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   629
fun add_desired_line _ _ _ _ (line as Definition (name, _, _)) (j, lines) =
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   630
    (j, line :: map (replace_dependencies_in_line (name, [])) lines)
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   631
  | add_desired_line isar_shrink_factor conjecture_shape fact_names frees
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   632
                     (Inference (name, t, rule, deps)) (j, lines) =
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   633
    (j + 1,
44773
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   634
     if is_fact fact_names name orelse
e701dabbfe37 perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents: 44399
diff changeset
   635
        is_conjecture conjecture_shape name orelse
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   636
        (* the last line must be kept *)
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   637
        j = 0 orelse
36570
9bebcb40599f identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
blanchet
parents: 36567
diff changeset
   638
        (not (is_only_type_information t) andalso
9bebcb40599f identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
blanchet
parents: 36567
diff changeset
   639
         null (Term.add_tvars t []) andalso
9bebcb40599f identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
blanchet
parents: 36567
diff changeset
   640
         not (exists_subterm (is_bad_free frees) t) andalso
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   641
         length deps >= 2 andalso j mod isar_shrink_factor = 0 andalso
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   642
         (* kill next to last line, which usually results in a trivial step *)
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   643
         j <> 1) then
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   644
       Inference (name, t, rule, deps) :: lines  (* keep line *)
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   645
     else
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   646
       map (replace_dependencies_in_line (name, deps)) lines)  (* drop line *)
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   647
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   648
(** Isar proof construction and manipulation **)
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   649
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   650
fun merge_fact_sets (ls1, ss1) (ls2, ss2) =
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   651
  (union (op =) ls1 ls2, union (op =) ss1 ss2)
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   652
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   653
type label = string * int
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   654
type facts = label list * string list
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   655
39452
70a57e40f795 factored out TSTP/SPASS/Vampire proof parsing;
blanchet
parents: 39425
diff changeset
   656
datatype isar_qualifier = Show | Then | Moreover | Ultimately
36291
b4c2043cc96c added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
blanchet
parents: 36288
diff changeset
   657
39452
70a57e40f795 factored out TSTP/SPASS/Vampire proof parsing;
blanchet
parents: 39425
diff changeset
   658
datatype isar_step =
36478
1aba777a367f fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents: 36477
diff changeset
   659
  Fix of (string * typ) list |
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   660
  Let of term * term |
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   661
  Assume of label * term |
39452
70a57e40f795 factored out TSTP/SPASS/Vampire proof parsing;
blanchet
parents: 39425
diff changeset
   662
  Have of isar_qualifier list * label * term * byline
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   663
and byline =
36564
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
   664
  ByMetis of facts |
39452
70a57e40f795 factored out TSTP/SPASS/Vampire proof parsing;
blanchet
parents: 39425
diff changeset
   665
  CaseSplit of isar_step list list * facts
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   666
36574
870dfa6d00ce eliminate trivial case splits from Isar proofs
blanchet
parents: 36570
diff changeset
   667
fun smart_case_split [] facts = ByMetis facts
870dfa6d00ce eliminate trivial case splits from Isar proofs
blanchet
parents: 36570
diff changeset
   668
  | smart_case_split proofs facts = CaseSplit (proofs, facts)
870dfa6d00ce eliminate trivial case splits from Isar proofs
blanchet
parents: 36570
diff changeset
   669
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   670
fun add_fact_from_dependency conjecture_shape facts_offset fact_names name =
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   671
  if is_fact fact_names name then
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   672
    apsnd (union (op =) (map fst (resolve_fact facts_offset fact_names name)))
36475
05209b869a6b new Isar proof construction code: stringfy axiom names correctly
blanchet
parents: 36474
diff changeset
   673
  else
39370
f8292d3020db use same hack as in "Async_Manager" to work around Proof General bug
blanchet
parents: 39368
diff changeset
   674
    apfst (insert (op =) (raw_label_for_name conjecture_shape name))
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   675
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   676
fun step_for_line _ _ _ _ (Definition (_, t1, t2)) = Let (t1, t2)
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   677
  | step_for_line conjecture_shape _ _ _ (Inference (name, t, _, [])) =
39370
f8292d3020db use same hack as in "Async_Manager" to work around Proof General bug
blanchet
parents: 39368
diff changeset
   678
    Assume (raw_label_for_name conjecture_shape name, t)
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   679
  | step_for_line conjecture_shape facts_offset fact_names j
45209
0e5e56e32bc0 cleaner LEO-II extensionality step detection
blanchet
parents: 45042
diff changeset
   680
                  (Inference (name, t, _, deps)) =
39370
f8292d3020db use same hack as in "Async_Manager" to work around Proof General bug
blanchet
parents: 39368
diff changeset
   681
    Have (if j = 1 then [Show] else [],
39425
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   682
          raw_label_for_name conjecture_shape name,
5d97fd83ab37 fix parsing of higher-order formulas;
blanchet
parents: 39374
diff changeset
   683
          fold_rev forall_of (map Var (Term.add_vars t [])) t,
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   684
          ByMetis (fold (add_fact_from_dependency conjecture_shape facts_offset
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   685
                                                  fact_names)
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   686
                        deps ([], [])))
36291
b4c2043cc96c added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
blanchet
parents: 36288
diff changeset
   687
39454
acb25e9cf6fb factor out the inverse of "nice_atp_problem"
blanchet
parents: 39453
diff changeset
   688
fun repair_name "$true" = "c_True"
acb25e9cf6fb factor out the inverse of "nice_atp_problem"
blanchet
parents: 39453
diff changeset
   689
  | repair_name "$false" = "c_False"
43000
bd424c3dde46 cleaner handling of equality and proxies (esp. for THF)
blanchet
parents: 42998
diff changeset
   690
  | repair_name "$$e" = tptp_equal (* seen in Vampire proofs *)
39454
acb25e9cf6fb factor out the inverse of "nice_atp_problem"
blanchet
parents: 39453
diff changeset
   691
  | repair_name s =
43000
bd424c3dde46 cleaner handling of equality and proxies (esp. for THF)
blanchet
parents: 42998
diff changeset
   692
    if is_tptp_equal s orelse
bd424c3dde46 cleaner handling of equality and proxies (esp. for THF)
blanchet
parents: 42998
diff changeset
   693
       (* seen in Vampire proofs *)
bd424c3dde46 cleaner handling of equality and proxies (esp. for THF)
blanchet
parents: 42998
diff changeset
   694
       (String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then
bd424c3dde46 cleaner handling of equality and proxies (esp. for THF)
blanchet
parents: 42998
diff changeset
   695
      tptp_equal
39454
acb25e9cf6fb factor out the inverse of "nice_atp_problem"
blanchet
parents: 39453
diff changeset
   696
    else
acb25e9cf6fb factor out the inverse of "nice_atp_problem"
blanchet
parents: 39453
diff changeset
   697
      s
acb25e9cf6fb factor out the inverse of "nice_atp_problem"
blanchet
parents: 39453
diff changeset
   698
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   699
fun isar_proof_from_atp_proof pool ctxt isar_shrink_factor conjecture_shape
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   700
        facts_offset fact_names sym_tab params frees atp_proof =
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   701
  let
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   702
    val lines =
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42361
diff changeset
   703
      atp_proof
42968
74415622d293 more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents: 42966
diff changeset
   704
      |> clean_up_atp_proof_dependencies
39454
acb25e9cf6fb factor out the inverse of "nice_atp_problem"
blanchet
parents: 39453
diff changeset
   705
      |> nasty_atp_proof pool
acb25e9cf6fb factor out the inverse of "nice_atp_problem"
blanchet
parents: 39453
diff changeset
   706
      |> map_term_names_in_atp_proof repair_name
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
   707
      |> decode_lines ctxt sym_tab
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   708
      |> rpair [] |-> fold_rev (add_line conjecture_shape fact_names)
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   709
      |> rpair [] |-> fold_rev add_nontrivial_line
42647
59142dbfa3ba no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
blanchet
parents: 42613
diff changeset
   710
      |> rpair (0, [])
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   711
      |-> fold_rev (add_desired_line isar_shrink_factor conjecture_shape
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   712
                                     fact_names frees)
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   713
      |> snd
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   714
  in
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
   715
    (if null params then [] else [Fix params]) @
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
   716
    map2 (step_for_line conjecture_shape facts_offset fact_names)
42541
8938507b2054 move type declarations to the front, for TFF-compliance
blanchet
parents: 42539
diff changeset
   717
         (length lines downto 1) lines
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   718
  end
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   719
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   720
(* When redirecting proofs, we keep information about the labels seen so far in
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   721
   the "backpatches" data structure. The first component indicates which facts
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   722
   should be associated with forthcoming proof steps. The second component is a
37322
0347b1a16682 fix bug in direct Isar proofs, which was exhibited by the "BigO" example
blanchet
parents: 37319
diff changeset
   723
   pair ("assum_ls", "drop_ls"), where "assum_ls" are the labels that should
0347b1a16682 fix bug in direct Isar proofs, which was exhibited by the "BigO" example
blanchet
parents: 37319
diff changeset
   724
   become assumptions and "drop_ls" are the labels that should be dropped in a
0347b1a16682 fix bug in direct Isar proofs, which was exhibited by the "BigO" example
blanchet
parents: 37319
diff changeset
   725
   case split. *)
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   726
type backpatches = (label * facts) list * (label list * label list)
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   727
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36555
diff changeset
   728
fun used_labels_of_step (Have (_, _, _, by)) =
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   729
    (case by of
36564
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
   730
       ByMetis (ls, _) => ls
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36555
diff changeset
   731
     | CaseSplit (proofs, (ls, _)) =>
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36555
diff changeset
   732
       fold (union (op =) o used_labels_of) proofs ls)
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36555
diff changeset
   733
  | used_labels_of_step _ = []
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36555
diff changeset
   734
and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof []
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   735
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   736
fun new_labels_of_step (Fix _) = []
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   737
  | new_labels_of_step (Let _) = []
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   738
  | new_labels_of_step (Assume (l, _)) = [l]
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   739
  | new_labels_of_step (Have (_, l, _, _)) = [l]
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   740
val new_labels_of = maps new_labels_of_step
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   741
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   742
val join_proofs =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   743
  let
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   744
    fun aux _ [] = NONE
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   745
      | aux proof_tail (proofs as (proof1 :: _)) =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   746
        if exists null proofs then
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   747
          NONE
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   748
        else if forall (curry (op =) (hd proof1) o hd) (tl proofs) then
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   749
          aux (hd proof1 :: proof_tail) (map tl proofs)
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   750
        else case hd proof1 of
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37479
diff changeset
   751
          Have ([], l, t, _) => (* FIXME: should we really ignore the "by"? *)
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   752
          if forall (fn Have ([], l', t', _) :: _ => (l, t) = (l', t')
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   753
                      | _ => false) (tl proofs) andalso
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   754
             not (exists (member (op =) (maps new_labels_of proofs))
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36555
diff changeset
   755
                         (used_labels_of proof_tail)) then
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   756
            SOME (l, t, map rev proofs, proof_tail)
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   757
          else
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   758
            NONE
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   759
        | _ => NONE
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   760
  in aux [] o map rev end
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   761
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   762
fun case_split_qualifiers proofs =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   763
  case length proofs of
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   764
    0 => []
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   765
  | 1 => [Then]
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   766
  | _ => [Ultimately]
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   767
39372
2fd8a9a7080d first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
blanchet
parents: 39370
diff changeset
   768
fun redirect_proof hyp_ts concl_t proof =
33310
44f9665c2091 proper header;
wenzelm
parents: 33243
diff changeset
   769
  let
37324
d77250dd2416 fix bugs in Sledgehammer's Isar proof "redirection" code
blanchet
parents: 37323
diff changeset
   770
    (* The first pass outputs those steps that are independent of the negated
d77250dd2416 fix bugs in Sledgehammer's Isar proof "redirection" code
blanchet
parents: 37323
diff changeset
   771
       conjecture. The second pass flips the proof by contradiction to obtain a
d77250dd2416 fix bugs in Sledgehammer's Isar proof "redirection" code
blanchet
parents: 37323
diff changeset
   772
       direct proof, introducing case splits when an inference depends on
d77250dd2416 fix bugs in Sledgehammer's Isar proof "redirection" code
blanchet
parents: 37323
diff changeset
   773
       several facts that depend on the negated conjecture. *)
39372
2fd8a9a7080d first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
blanchet
parents: 39370
diff changeset
   774
     val concl_l = (conjecture_prefix, length hyp_ts)
38040
174568533593 fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents: 38039
diff changeset
   775
     fun first_pass ([], contra) = ([], contra)
174568533593 fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents: 38039
diff changeset
   776
       | first_pass ((step as Fix _) :: proof, contra) =
174568533593 fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents: 38039
diff changeset
   777
         first_pass (proof, contra) |>> cons step
174568533593 fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents: 38039
diff changeset
   778
       | first_pass ((step as Let _) :: proof, contra) =
174568533593 fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents: 38039
diff changeset
   779
         first_pass (proof, contra) |>> cons step
39370
f8292d3020db use same hack as in "Async_Manager" to work around Proof General bug
blanchet
parents: 39368
diff changeset
   780
       | first_pass ((step as Assume (l as (_, j), _)) :: proof, contra) =
39372
2fd8a9a7080d first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
blanchet
parents: 39370
diff changeset
   781
         if l = concl_l then first_pass (proof, contra ||> cons step)
2fd8a9a7080d first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
blanchet
parents: 39370
diff changeset
   782
         else first_pass (proof, contra) |>> cons (Assume (l, nth hyp_ts j))
38040
174568533593 fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents: 38039
diff changeset
   783
       | first_pass (Have (qs, l, t, ByMetis (ls, ss)) :: proof, contra) =
39372
2fd8a9a7080d first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
blanchet
parents: 39370
diff changeset
   784
         let val step = Have (qs, l, t, ByMetis (ls, ss)) in
38040
174568533593 fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents: 38039
diff changeset
   785
           if exists (member (op =) (fst contra)) ls then
174568533593 fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents: 38039
diff changeset
   786
             first_pass (proof, contra |>> cons l ||> cons step)
174568533593 fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents: 38039
diff changeset
   787
           else
174568533593 fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents: 38039
diff changeset
   788
             first_pass (proof, contra) |>> cons step
174568533593 fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents: 38039
diff changeset
   789
         end
174568533593 fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
blanchet
parents: 38039
diff changeset
   790
       | first_pass _ = raise Fail "malformed proof"
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   791
    val (proof_top, (contra_ls, contra_proof)) =
39372
2fd8a9a7080d first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
blanchet
parents: 39370
diff changeset
   792
      first_pass (proof, ([concl_l], []))
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   793
    val backpatch_label = the_default ([], []) oo AList.lookup (op =) o fst
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   794
    fun backpatch_labels patches ls =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   795
      fold merge_fact_sets (map (backpatch_label patches) ls) ([], [])
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   796
    fun second_pass end_qs ([], assums, patches) =
37324
d77250dd2416 fix bugs in Sledgehammer's Isar proof "redirection" code
blanchet
parents: 37323
diff changeset
   797
        ([Have (end_qs, no_label, concl_t,
36564
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
   798
                ByMetis (backpatch_labels patches (map snd assums)))], patches)
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   799
      | second_pass end_qs (Assume (l, t) :: proof, assums, patches) =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   800
        second_pass end_qs (proof, (t, l) :: assums, patches)
36564
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
   801
      | second_pass end_qs (Have (qs, l, t, ByMetis (ls, ss)) :: proof, assums,
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   802
                            patches) =
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   803
        (if member (op =) (snd (snd patches)) l andalso
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   804
            not (member (op =) (fst (snd patches)) l) andalso
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   805
            not (AList.defined (op =) (fst patches) l) then
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   806
           second_pass end_qs (proof, assums, patches ||> apsnd (append ls))
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   807
         else case List.partition (member (op =) contra_ls) ls of
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   808
           ([contra_l], co_ls) =>
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   809
           if member (op =) qs Show then
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   810
             second_pass end_qs (proof, assums,
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   811
                                 patches |>> cons (contra_l, (co_ls, ss)))
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   812
           else
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   813
             second_pass end_qs
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   814
                         (proof, assums,
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   815
                          patches |>> cons (contra_l, (l :: co_ls, ss)))
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   816
             |>> cons (if member (op =) (fst (snd patches)) l then
42606
0c76cf483899 show sorts not just types in Isar proofs + tuning
blanchet
parents: 42595
diff changeset
   817
                         Assume (l, s_not t)
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   818
                       else
42606
0c76cf483899 show sorts not just types in Isar proofs + tuning
blanchet
parents: 42595
diff changeset
   819
                         Have (qs, l, s_not t,
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   820
                               ByMetis (backpatch_label patches l)))
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   821
         | (contra_ls as _ :: _, co_ls) =>
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   822
           let
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   823
             val proofs =
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   824
               map_filter
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   825
                   (fn l =>
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   826
                       if l = concl_l then
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   827
                         NONE
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   828
                       else
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   829
                         let
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   830
                           val drop_ls = filter (curry (op <>) l) contra_ls
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   831
                         in
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   832
                           second_pass []
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   833
                               (proof, assums,
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   834
                                patches ||> apfst (insert (op =) l)
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   835
                                        ||> apsnd (union (op =) drop_ls))
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   836
                           |> fst |> SOME
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   837
                         end) contra_ls
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   838
             val (assumes, facts) =
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   839
               if member (op =) (fst (snd patches)) l then
42606
0c76cf483899 show sorts not just types in Isar proofs + tuning
blanchet
parents: 42595
diff changeset
   840
                 ([Assume (l, s_not t)], (l :: co_ls, ss))
39373
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   841
               else
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   842
                 ([], (co_ls, ss))
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   843
           in
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   844
             (case join_proofs proofs of
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   845
                SOME (l, t, proofs, proof_tail) =>
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   846
                Have (case_split_qualifiers proofs @
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   847
                      (if null proof_tail then end_qs else []), l, t,
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   848
                      smart_case_split proofs facts) :: proof_tail
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   849
              | NONE =>
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   850
                [Have (case_split_qualifiers proofs @ end_qs, no_label,
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   851
                       concl_t, smart_case_split proofs facts)],
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   852
              patches)
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   853
             |>> append assumes
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   854
           end
fe95c860434c finish support for E 1.2 proof reconstruction;
blanchet
parents: 39372
diff changeset
   855
         | _ => raise Fail "malformed proof")
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   856
       | second_pass _ _ = raise Fail "malformed proof"
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   857
    val proof_bottom =
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   858
      second_pass [Show] (contra_proof, [], ([], ([], []))) |> fst
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   859
  in proof_top @ proof_bottom end
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   860
38490
blanchet
parents: 38488
diff changeset
   861
(* FIXME: Still needed? Probably not. *)
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   862
val kill_duplicate_assumptions_in_proof =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   863
  let
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   864
    fun relabel_facts subst =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   865
      apfst (map (fn l => AList.lookup (op =) subst l |> the_default l))
36491
6769f8eacaac unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
blanchet
parents: 36488
diff changeset
   866
    fun do_step (step as Assume (l, t)) (proof, subst, assums) =
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   867
        (case AList.lookup (op aconv) assums t of
36967
3c804030474b fix bug in Isar proof reconstruction step relabeling + don't try to infer the sorts of TVars, since this often fails miserably
blanchet
parents: 36924
diff changeset
   868
           SOME l' => (proof, (l, l') :: subst, assums)
36491
6769f8eacaac unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
blanchet
parents: 36488
diff changeset
   869
         | NONE => (step :: proof, subst, (t, l) :: assums))
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   870
      | do_step (Have (qs, l, t, by)) (proof, subst, assums) =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   871
        (Have (qs, l, t,
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   872
               case by of
36564
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
   873
                 ByMetis facts => ByMetis (relabel_facts subst facts)
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   874
               | CaseSplit (proofs, facts) =>
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   875
                 CaseSplit (map do_proof proofs, relabel_facts subst facts)) ::
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   876
         proof, subst, assums)
36491
6769f8eacaac unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
blanchet
parents: 36488
diff changeset
   877
      | do_step step (proof, subst, assums) = (step :: proof, subst, assums)
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   878
    and do_proof proof = fold do_step proof ([], [], []) |> #1 |> rev
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   879
  in do_proof end
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   880
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   881
val then_chain_proof =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   882
  let
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   883
    fun aux _ [] = []
36491
6769f8eacaac unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
blanchet
parents: 36488
diff changeset
   884
      | aux _ ((step as Assume (l, _)) :: proof) = step :: aux l proof
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   885
      | aux l' (Have (qs, l, t, by) :: proof) =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   886
        (case by of
36564
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
   887
           ByMetis (ls, ss) =>
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   888
           Have (if member (op =) ls l' then
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   889
                   (Then :: qs, l, t,
36564
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
   890
                    ByMetis (filter_out (curry (op =) l') ls, ss))
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   891
                 else
36564
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
   892
                   (qs, l, t, ByMetis (ls, ss)))
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   893
         | CaseSplit (proofs, facts) =>
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   894
           Have (qs, l, t, CaseSplit (map (aux no_label) proofs, facts))) ::
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   895
        aux l proof
36491
6769f8eacaac unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
blanchet
parents: 36488
diff changeset
   896
      | aux _ (step :: proof) = step :: aux no_label proof
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   897
  in aux no_label end
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   898
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   899
fun kill_useless_labels_in_proof proof =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   900
  let
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36555
diff changeset
   901
    val used_ls = used_labels_of proof
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   902
    fun do_label l = if member (op =) used_ls l then l else no_label
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36555
diff changeset
   903
    fun do_step (Assume (l, t)) = Assume (do_label l, t)
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36555
diff changeset
   904
      | do_step (Have (qs, l, t, by)) =
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   905
        Have (qs, do_label l, t,
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   906
              case by of
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   907
                CaseSplit (proofs, facts) =>
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36555
diff changeset
   908
                CaseSplit (map (map do_step) proofs, facts)
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   909
              | _ => by)
36556
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36555
diff changeset
   910
      | do_step step = step
81dc2c20f052 use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents: 36555
diff changeset
   911
  in map do_step proof end
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   912
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   913
fun prefix_for_depth n = replicate_string (n + 1)
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   914
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   915
val relabel_proof =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   916
  let
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   917
    fun aux _ _ _ [] = []
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   918
      | aux subst depth (next_assum, next_fact) (Assume (l, t) :: proof) =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   919
        if l = no_label then
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   920
          Assume (l, t) :: aux subst depth (next_assum, next_fact) proof
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   921
        else
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   922
          let val l' = (prefix_for_depth depth assum_prefix, next_assum) in
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   923
            Assume (l', t) ::
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   924
            aux ((l, l') :: subst) depth (next_assum + 1, next_fact) proof
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   925
          end
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   926
      | aux subst depth (next_assum, next_fact) (Have (qs, l, t, by) :: proof) =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   927
        let
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   928
          val (l', subst, next_fact) =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   929
            if l = no_label then
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   930
              (l, subst, next_fact)
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   931
            else
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   932
              let
42180
a6c141925a8a added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents: 41742
diff changeset
   933
                val l' = (prefix_for_depth depth have_prefix, next_fact)
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   934
              in (l', (l, l') :: subst, next_fact + 1) end
36570
9bebcb40599f identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
blanchet
parents: 36567
diff changeset
   935
          val relabel_facts =
39370
f8292d3020db use same hack as in "Async_Manager" to work around Proof General bug
blanchet
parents: 39368
diff changeset
   936
            apfst (maps (the_list o AList.lookup (op =) subst))
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   937
          val by =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   938
            case by of
36564
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
   939
              ByMetis facts => ByMetis (relabel_facts facts)
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   940
            | CaseSplit (proofs, facts) =>
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   941
              CaseSplit (map (aux subst (depth + 1) (1, 1)) proofs,
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   942
                         relabel_facts facts)
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   943
        in
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   944
          Have (qs, l', t, by) ::
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   945
          aux subst depth (next_assum, next_fact) proof
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   946
        end
36491
6769f8eacaac unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
blanchet
parents: 36488
diff changeset
   947
      | aux subst depth nextp (step :: proof) =
6769f8eacaac unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
blanchet
parents: 36488
diff changeset
   948
        step :: aux subst depth nextp proof
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   949
  in aux [] 0 (1, 1) end
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   950
42881
dbdfe2d5b6ab automatically use "metisFT" when typed helpers are necessary
blanchet
parents: 42876
diff changeset
   951
fun string_for_proof ctxt0 full_types i n =
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   952
  let
42761
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   953
    val ctxt =
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   954
      ctxt0 |> Config.put show_free_types false
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   955
            |> Config.put show_types true
8ea9c6fa8b53 fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
blanchet
parents: 42759
diff changeset
   956
            |> Config.put show_sorts true
37319
42268dc7d6c4 show types in Isar proofs, but not for free variables;
blanchet
parents: 37172
diff changeset
   957
    fun fix_print_mode f x =
39134
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 39115
diff changeset
   958
      Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN)
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 39115
diff changeset
   959
                               (print_mode_value ())) f x
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   960
    fun do_indent ind = replicate_string (ind * indent_size) " "
36478
1aba777a367f fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents: 36477
diff changeset
   961
    fun do_free (s, T) =
1aba777a367f fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents: 36477
diff changeset
   962
      maybe_quote s ^ " :: " ^
1aba777a367f fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents: 36477
diff changeset
   963
      maybe_quote (fix_print_mode (Syntax.string_of_typ ctxt) T)
36570
9bebcb40599f identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
blanchet
parents: 36567
diff changeset
   964
    fun do_label l = if l = no_label then "" else string_for_label l ^ ": "
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   965
    fun do_have qs =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   966
      (if member (op =) qs Moreover then "moreover " else "") ^
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   967
      (if member (op =) qs Ultimately then "ultimately " else "") ^
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   968
      (if member (op =) qs Then then
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   969
         if member (op =) qs Show then "thus" else "hence"
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   970
       else
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   971
         if member (op =) qs Show then "show" else "have")
36478
1aba777a367f fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents: 36477
diff changeset
   972
    val do_term = maybe_quote o fix_print_mode (Syntax.string_of_term ctxt)
43228
2ed2f092e990 obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents: 43212
diff changeset
   973
    val reconstructor = if full_types then Metis_Full_Types else Metis
36570
9bebcb40599f identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
blanchet
parents: 36567
diff changeset
   974
    fun do_facts (ls, ss) =
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   975
      reconstructor_command reconstructor 1 1
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   976
          (ls |> sort_distinct (prod_ord string_ord int_ord),
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
   977
           ss |> sort_distinct string_ord)
36478
1aba777a367f fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents: 36477
diff changeset
   978
    and do_step ind (Fix xs) =
1aba777a367f fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents: 36477
diff changeset
   979
        do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n"
36486
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   980
      | do_step ind (Let (t1, t2)) =
c2d7e2dff59e support Vampire definitions of constants as "let" constructs in Isar proofs
blanchet
parents: 36485
diff changeset
   981
        do_indent ind ^ "let " ^ do_term t1 ^ " = " ^ do_term t2 ^ "\n"
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   982
      | do_step ind (Assume (l, t)) =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   983
        do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n"
36564
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
   984
      | do_step ind (Have (qs, l, t, ByMetis facts)) =
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   985
        do_indent ind ^ do_have qs ^ " " ^
36479
fcbf412c560f reintroduce missing "gen_all_vars" call
blanchet
parents: 36478
diff changeset
   986
        do_label l ^ do_term t ^ " " ^ do_facts facts ^ "\n"
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   987
      | do_step ind (Have (qs, l, t, CaseSplit (proofs, facts))) =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   988
        space_implode (do_indent ind ^ "moreover\n")
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   989
                      (map (do_block ind) proofs) ^
36479
fcbf412c560f reintroduce missing "gen_all_vars" call
blanchet
parents: 36478
diff changeset
   990
        do_indent ind ^ do_have qs ^ " " ^ do_label l ^ do_term t ^ " " ^
36478
1aba777a367f fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents: 36477
diff changeset
   991
        do_facts facts ^ "\n"
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   992
    and do_steps prefix suffix ind steps =
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   993
      let val s = implode (map (do_step ind) steps) in
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   994
        replicate_string (ind * indent_size - size prefix) " " ^ prefix ^
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   995
        String.extract (s, ind * indent_size,
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   996
                        SOME (size s - ind * indent_size - 1)) ^
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   997
        suffix ^ "\n"
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   998
      end
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
   999
    and do_block ind proof = do_steps "{ " " }" (ind + 1) proof
36564
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
  1000
    (* One-step proofs are pointless; better use the Metis one-liner
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
  1001
       directly. *)
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
  1002
    and do_proof [Have (_, _, _, ByMetis _)] = ""
96f767f546e7 be more discriminate: some one-line Isar proofs are silly
blanchet
parents: 36563
diff changeset
  1003
      | do_proof proof =
36480
1e01a7162435 polish Isar proofs: don't mention facts twice, and don't show one-liner "structured" proofs
blanchet
parents: 36479
diff changeset
  1004
        (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^
39452
70a57e40f795 factored out TSTP/SPASS/Vampire proof parsing;
blanchet
parents: 39425
diff changeset
  1005
        do_indent 0 ^ "proof -\n" ^ do_steps "" "" 1 proof ^ do_indent 0 ^
70a57e40f795 factored out TSTP/SPASS/Vampire proof parsing;
blanchet
parents: 39425
diff changeset
  1006
        (if n <> 1 then "next" else "qed")
36488
32c92af68ec9 remove Sledgehammer's "sorts" option to annotate variables with sorts in proof;
blanchet
parents: 36486
diff changeset
  1007
  in do_proof end
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
  1008
43062
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1009
fun isar_proof_text ctxt isar_proof_requested
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
  1010
        (debug, full_types, isar_shrink_factor, pool, conjecture_shape,
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
  1011
         facts_offset, fact_names, sym_tab, atp_proof, goal)
43037
ade5c84f860f cleanup proof text generation code
blanchet
parents: 43034
diff changeset
  1012
        (one_line_params as (_, _, _, _, subgoal, subgoal_count)) =
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
  1013
  let
43062
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1014
    val isar_shrink_factor =
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1015
      (if isar_proof_requested then 1 else 2) * isar_shrink_factor
43037
ade5c84f860f cleanup proof text generation code
blanchet
parents: 43034
diff changeset
  1016
    val (params, hyp_ts, concl_t) = strip_subgoal ctxt goal subgoal
36909
7d5587f6d5f7 made Sledgehammer's full-typed proof reconstruction work for the first time;
blanchet
parents: 36607
diff changeset
  1017
    val frees = fold Term.add_frees (concl_t :: hyp_ts) []
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
  1018
    val one_line_proof = one_line_proof_text one_line_params
36283
25e69e93954d failure of reconstructing the Isar proof (e.g., exception) should not prevent Sledgehammer from showing the short one-liner proof -- but in "debug" mode we do want to know what the exception is
blanchet
parents: 36281
diff changeset
  1019
    fun isar_proof_for () =
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
  1020
      case atp_proof
43304
6901ebafbb8d cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
blanchet
parents: 43302
diff changeset
  1021
           |> isar_proof_from_atp_proof pool ctxt isar_shrink_factor
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
  1022
                  conjecture_shape facts_offset fact_names sym_tab params frees
39372
2fd8a9a7080d first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
blanchet
parents: 39370
diff changeset
  1023
           |> redirect_proof hyp_ts concl_t
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
  1024
           |> kill_duplicate_assumptions_in_proof
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
  1025
           |> then_chain_proof
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
  1026
           |> kill_useless_labels_in_proof
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
  1027
           |> relabel_proof
43037
ade5c84f860f cleanup proof text generation code
blanchet
parents: 43034
diff changeset
  1028
           |> string_for_proof ctxt full_types subgoal subgoal_count of
43062
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1029
        "" =>
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1030
        if isar_proof_requested then
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1031
          "\nNo structured proof available (proof too short)."
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1032
        else
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1033
          ""
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1034
      | proof =>
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1035
        "\n\n" ^ (if isar_proof_requested then "Structured proof"
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1036
                  else "Perhaps this will work") ^
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1037
        ":\n" ^ Markup.markup Markup.sendback proof
35868
491a97039ce1 renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents: 35865
diff changeset
  1038
    val isar_proof =
36402
1b20805974c7 introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents: 36396
diff changeset
  1039
      if debug then
36283
25e69e93954d failure of reconstructing the Isar proof (e.g., exception) should not prevent Sledgehammer from showing the short one-liner proof -- but in "debug" mode we do want to know what the exception is
blanchet
parents: 36281
diff changeset
  1040
        isar_proof_for ()
25e69e93954d failure of reconstructing the Isar proof (e.g., exception) should not prevent Sledgehammer from showing the short one-liner proof -- but in "debug" mode we do want to know what the exception is
blanchet
parents: 36281
diff changeset
  1041
      else
43062
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1042
        case try isar_proof_for () of
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1043
          SOME s => s
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1044
        | NONE => if isar_proof_requested then
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1045
                    "\nWarning: The Isar proof construction failed."
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1046
                  else
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1047
                    ""
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
  1048
  in one_line_proof ^ isar_proof end
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
  1049
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
  1050
fun proof_text ctxt isar_proof isar_params
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
  1051
               (one_line_params as (preplay, _, _, _, _, _)) =
43166
68e3cd19fee8 show what failed to play
blanchet
parents: 43163
diff changeset
  1052
  (if case preplay of Failed_to_Play _ => true | _ => isar_proof then
43062
2834548a7a48 nicer failure message when one-line proof reconstruction fails
blanchet
parents: 43051
diff changeset
  1053
     isar_proof_text ctxt isar_proof isar_params
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
  1054
   else
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43006
diff changeset
  1055
     one_line_proof_text) one_line_params
36223
217ca1273786 make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents: 36140
diff changeset
  1056
31038
immler@in.tum.de
parents: 31037
diff changeset
  1057
end;