src/HOL/Tools/ATP/atp_proof.ML
author blanchet
Thu, 24 Jul 2014 18:46:38 +0200
changeset 58998 49077e289606
parent 58635 4e619ee65a61
child 59180 85ec71012df8
permissions -rw-r--r--
filter out 'theory(...)' from dependencies early on
blanchet@39692
     1
(*  Title:      HOL/Tools/ATP/atp_proof.ML
blanchet@39692
     2
    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
blanchet@39692
     3
    Author:     Claire Quigley, Cambridge University Computer Laboratory
blanchet@39692
     4
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@39692
     5
blanchet@43745
     6
Abstract representation of ATP proofs and TSTP/SPASS syntax.
blanchet@39692
     7
*)
blanchet@39692
     8
blanchet@39692
     9
signature ATP_PROOF =
blanchet@39692
    10
sig
blanchet@56153
    11
  type 'a atp_type = 'a ATP_Problem.atp_type
blanchet@54723
    12
  type ('a, 'b) atp_term = ('a, 'b) ATP_Problem.atp_term
blanchet@54723
    13
  type atp_formula_role = ATP_Problem.atp_formula_role
blanchet@54723
    14
  type ('a, 'b, 'c, 'd) atp_formula = ('a, 'b, 'c, 'd) ATP_Problem.atp_formula
blanchet@54723
    15
  type 'a atp_problem = 'a ATP_Problem.atp_problem
blanchet@39692
    16
blanchet@43806
    17
  exception UNRECOGNIZED_ATP_PROOF of unit
blanchet@43806
    18
blanchet@54723
    19
  datatype atp_failure =
blanchet@43458
    20
    Unprovable |
blanchet@43891
    21
    GaveUp |
blanchet@43458
    22
    ProofMissing |
blanchet@43751
    23
    ProofIncomplete |
blanchet@58608
    24
    ProofUnparsable |
blanchet@45786
    25
    UnsoundProof of bool * string list |
blanchet@43458
    26
    CantConnect |
blanchet@43458
    27
    TimedOut |
blanchet@43794
    28
    Inappropriate |
blanchet@43458
    29
    OutOfResources |
blanchet@43458
    30
    NoPerl |
blanchet@43458
    31
    NoLibwwwPerl |
blanchet@43458
    32
    MalformedInput |
blanchet@43458
    33
    MalformedOutput |
blanchet@43458
    34
    Interrupted |
blanchet@43458
    35
    Crashed |
blanchet@43458
    36
    InternalError |
blanchet@43458
    37
    UnknownError of string
blanchet@39731
    38
blanchet@54723
    39
  type atp_step_name = string * string list
blanchet@54724
    40
  type ('a, 'b) atp_step =
blanchet@54724
    41
    atp_step_name * atp_formula_role * 'a * 'b * atp_step_name list
blanchet@39692
    42
blanchet@56153
    43
  type 'a atp_proof = (('a, 'a, ('a, 'a atp_type) atp_term, 'a) atp_formula, string) atp_step list
blanchet@39692
    44
fleury@58496
    45
  (* Named ATPs *)
fleury@58496
    46
  val agsyholN : string
fleury@58496
    47
  val alt_ergoN : string
fleury@58496
    48
  val dummy_thfN : string
blanchet@58635
    49
  val dummy_thf_mlN : string
fleury@58496
    50
  val eN : string
fleury@58496
    51
  val e_malesN : string
fleury@58496
    52
  val e_parN : string
fleury@58496
    53
  val e_sineN : string
fleury@58496
    54
  val e_tofofN : string
fleury@58496
    55
  val iproverN : string
fleury@58496
    56
  val iprover_eqN : string
fleury@58496
    57
  val leo2N : string
fleury@58496
    58
  val satallaxN : string
fleury@58496
    59
  val snarkN : string
fleury@58496
    60
  val spassN : string
fleury@58496
    61
  val spass_pirateN : string
fleury@58496
    62
  val vampireN : string
fleury@58496
    63
  val waldmeisterN : string
blanchet@58557
    64
  val waldmeister_newN : string
fleury@58496
    65
  val z3_tptpN : string
fleury@58496
    66
  val zipperpositionN : string
fleury@58496
    67
  val remote_prefix : string
fleury@58496
    68
blanchet@56130
    69
  val agsyhol_core_rule : string
blanchet@56130
    70
  val satallax_core_rule : string
blanchet@56130
    71
  val spass_input_rule : string
blanchet@56534
    72
  val spass_pre_skolemize_rule : string
blanchet@56130
    73
  val spass_skolemize_rule : string
blanchet@57746
    74
  val z3_tptp_core_rule : string
blanchet@56130
    75
blanchet@41505
    76
  val short_output : bool -> string -> string
blanchet@54723
    77
  val string_of_atp_failure : atp_failure -> string
blanchet@39731
    78
  val extract_important_message : string -> string
blanchet@56153
    79
  val extract_known_atp_failure : (atp_failure * string) list -> string -> atp_failure option
blanchet@39731
    80
  val extract_tstplike_proof_and_outcome :
blanchet@54723
    81
    bool -> (string * string) list -> (atp_failure * string) list -> string
blanchet@54723
    82
    -> string * atp_failure option
blanchet@54723
    83
  val is_same_atp_step : atp_step_name -> atp_step_name -> bool
blanchet@43802
    84
  val scan_general_id : string list -> string * string list
blanchet@56153
    85
  val parse_formula : string list ->
blanchet@56153
    86
    (string, string atp_type, (string, string atp_type) atp_term, string) atp_formula * string list
fleury@58496
    87
  val atp_proof_of_tstplike_proof : string -> string atp_problem -> string -> string atp_proof
blanchet@54723
    88
  val clean_up_atp_proof_dependencies : string atp_proof -> string atp_proof
blanchet@56153
    89
  val map_term_names_in_atp_proof : (string -> string) -> string atp_proof -> string atp_proof
blanchet@56153
    90
  val nasty_atp_proof : string Symtab.table -> string atp_proof -> string atp_proof
blanchet@39692
    91
end;
blanchet@39692
    92
blanchet@39692
    93
structure ATP_Proof : ATP_PROOF =
blanchet@39692
    94
struct
blanchet@39692
    95
blanchet@43926
    96
open ATP_Util
blanchet@39731
    97
open ATP_Problem
blanchet@39731
    98
fleury@58496
    99
(* Named ATPs *)
fleury@58496
   100
fleury@58496
   101
val agsyholN = "agsyhol"
fleury@58496
   102
val alt_ergoN = "alt_ergo"
fleury@58496
   103
val dummy_thfN = "dummy_thf" (* for experiments *)
blanchet@58635
   104
val dummy_thf_mlN = "dummy_thf_ml" (* for experiments *)
fleury@58496
   105
val eN = "e"
fleury@58496
   106
val e_malesN = "e_males"
fleury@58496
   107
val e_parN = "e_par"
fleury@58496
   108
val e_sineN = "e_sine"
fleury@58496
   109
val e_tofofN = "e_tofof"
fleury@58496
   110
val iproverN = "iprover"
fleury@58496
   111
val iprover_eqN = "iprover_eq"
fleury@58496
   112
val leo2N = "leo2"
fleury@58496
   113
val satallaxN = "satallax"
fleury@58496
   114
val snarkN = "snark"
fleury@58496
   115
val spassN = "spass"
fleury@58496
   116
val spass_pirateN = "spass_pirate"
fleury@58496
   117
val vampireN = "vampire"
fleury@58496
   118
val waldmeisterN = "waldmeister"
blanchet@58557
   119
val waldmeister_newN = "waldmeister_new"
fleury@58496
   120
val z3_tptpN = "z3_tptp"
fleury@58496
   121
val zipperpositionN = "zipperposition"
fleury@58496
   122
val remote_prefix = "remote_"
fleury@58496
   123
blanchet@56130
   124
val agsyhol_core_rule = "__agsyhol_core" (* arbitrary *)
blanchet@56130
   125
val satallax_core_rule = "__satallax_core" (* arbitrary *)
blanchet@56130
   126
val spass_input_rule = "Inp"
blanchet@56534
   127
val spass_pre_skolemize_rule = "__Sko0" (* arbitrary *)
blanchet@56130
   128
val spass_skolemize_rule = "__Sko" (* arbitrary *)
blanchet@57746
   129
val z3_tptp_core_rule = "__z3_tptp_core" (* arbitrary *)
blanchet@56130
   130
blanchet@43806
   131
exception UNRECOGNIZED_ATP_PROOF of unit
blanchet@43806
   132
blanchet@54723
   133
datatype atp_failure =
blanchet@43458
   134
  Unprovable |
blanchet@43891
   135
  GaveUp |
blanchet@43458
   136
  ProofMissing |
blanchet@43751
   137
  ProofIncomplete |
blanchet@58608
   138
  ProofUnparsable |
blanchet@45786
   139
  UnsoundProof of bool * string list |
blanchet@43458
   140
  CantConnect |
blanchet@43458
   141
  TimedOut |
blanchet@43794
   142
  Inappropriate |
blanchet@43458
   143
  OutOfResources |
blanchet@43458
   144
  NoPerl |
blanchet@43458
   145
  NoLibwwwPerl |
blanchet@43458
   146
  MalformedInput |
blanchet@43458
   147
  MalformedOutput |
blanchet@43458
   148
  Interrupted |
blanchet@43458
   149
  Crashed |
blanchet@43458
   150
  InternalError |
blanchet@43458
   151
  UnknownError of string
blanchet@39731
   152
blanchet@41505
   153
fun short_output verbose output =
blanchet@42915
   154
  if verbose then
blanchet@42915
   155
    if output = "" then "No details available" else elide_string 1000 output
blanchet@42915
   156
  else
blanchet@42915
   157
    ""
blanchet@41505
   158
blanchet@42615
   159
val missing_message_tail =
blanchet@42615
   160
  " appears to be missing. You will need to install it if you want to invoke \
blanchet@42615
   161
  \remote provers."
blanchet@39731
   162
blanchet@56130
   163
fun from_lemmas [] = ""
blanchet@56130
   164
  | from_lemmas ss = " from " ^ space_implode " " (Try.serial_commas "and" (map quote ss))
blanchet@43745
   165
blanchet@54723
   166
fun string_of_atp_failure Unprovable = "The generated problem is unprovable."
blanchet@54723
   167
  | string_of_atp_failure GaveUp = "The prover gave up."
blanchet@54723
   168
  | string_of_atp_failure ProofMissing =
blanchet@42615
   169
    "The prover claims the conjecture is a theorem but did not provide a proof."
blanchet@54723
   170
  | string_of_atp_failure ProofIncomplete =
blanchet@58608
   171
    "The prover claims the conjecture is a theorem but provided an incomplete proof."
blanchet@58608
   172
  | string_of_atp_failure ProofUnparsable =
blanchet@58608
   173
    "The prover claims the conjecture is a theorem but provided an unparsable proof."
blanchet@54723
   174
  | string_of_atp_failure (UnsoundProof (false, ss)) =
blanchet@56130
   175
    "The prover derived \"False\"" ^ from_lemmas ss ^
blanchet@54361
   176
    ". Specify a sound type encoding or omit the \"type_enc\" option."
blanchet@54723
   177
  | string_of_atp_failure (UnsoundProof (true, ss)) =
blanchet@56130
   178
    "The prover derived \"False\"" ^ from_lemmas ss ^
blanchet@54361
   179
    ". This could be due to inconsistent axioms (including \"sorry\"s) or to \
blanchet@54361
   180
    \a bug in Sledgehammer. If the problem persists, please contact the \
blanchet@54361
   181
    \Isabelle developers."
blanchet@54723
   182
  | string_of_atp_failure CantConnect = "Cannot connect to remote server."
blanchet@54723
   183
  | string_of_atp_failure TimedOut = "Timed out."
blanchet@54723
   184
  | string_of_atp_failure Inappropriate =
blanchet@48377
   185
    "The generated problem lies outside the prover's scope."
blanchet@54723
   186
  | string_of_atp_failure OutOfResources = "The prover ran out of resources."
blanchet@54723
   187
  | string_of_atp_failure NoPerl = "Perl" ^ missing_message_tail
blanchet@54723
   188
  | string_of_atp_failure NoLibwwwPerl =
blanchet@42615
   189
    "The Perl module \"libwww-perl\"" ^ missing_message_tail
blanchet@54723
   190
  | string_of_atp_failure MalformedInput =
blanchet@42615
   191
    "The generated problem is malformed. Please report this to the Isabelle \
blanchet@42615
   192
    \developers."
blanchet@54723
   193
  | string_of_atp_failure MalformedOutput = "The prover output is malformed."
blanchet@54723
   194
  | string_of_atp_failure Interrupted = "The prover was interrupted."
blanchet@54723
   195
  | string_of_atp_failure Crashed = "The prover crashed."
blanchet@54723
   196
  | string_of_atp_failure InternalError = "An internal prover error occurred."
blanchet@54723
   197
  | string_of_atp_failure (UnknownError s) =
blanchet@42615
   198
    "A prover error occurred" ^
blanchet@53135
   199
    (if s = "" then ". (Pass the \"verbose\" option for details.)"
blanchet@53135
   200
     else ":\n" ^ s)
blanchet@39731
   201
blanchet@39731
   202
fun extract_delimited (begin_delim, end_delim) output =
blanchet@57739
   203
  (case first_field begin_delim output of
blanchet@57739
   204
    SOME (_, tail) =>
blanchet@57739
   205
    (case first_field "\n" tail of
blanchet@57739
   206
      SOME (_, tail') =>
blanchet@57739
   207
      if end_delim = "" then
blanchet@57739
   208
        tail'
blanchet@57739
   209
      else
blanchet@57739
   210
        (case first_field end_delim tail' of
blanchet@57739
   211
          SOME (body, _) => body
blanchet@57739
   212
        | NONE => "")
blanchet@57739
   213
    | NONE => "")
blanchet@57739
   214
  | NONE => "")
blanchet@39731
   215
blanchet@39731
   216
val tstp_important_message_delims =
blanchet@39731
   217
  ("% SZS start RequiredInformation", "% SZS end RequiredInformation")
blanchet@39731
   218
blanchet@39731
   219
fun extract_important_message output =
blanchet@56130
   220
  (case extract_delimited tstp_important_message_delims output of
blanchet@39731
   221
    "" => ""
blanchet@39731
   222
  | s => s |> space_explode "\n" |> filter_out (curry (op =) "")
blanchet@39731
   223
           |> map (perhaps (try (unprefix "%")))
blanchet@39731
   224
           |> map (perhaps (try (unprefix " ")))
blanchet@56130
   225
           |> space_implode "\n " |> quote)
blanchet@39731
   226
blanchet@39731
   227
(* Splits by the first possible of a list of delimiters. *)
blanchet@39731
   228
fun extract_tstplike_proof delims output =
blanchet@57739
   229
  (case pairself (find_first (fn s => String.isSubstring s output)) (ListPair.unzip delims) of
blanchet@56130
   230
    (SOME begin_delim, SOME end_delim) => extract_delimited (begin_delim, end_delim) output
blanchet@56130
   231
  | _ => "")
blanchet@39731
   232
blanchet@54723
   233
fun extract_known_atp_failure known_failures output =
blanchet@39731
   234
  known_failures
blanchet@39731
   235
  |> find_first (fn (_, pattern) => String.isSubstring pattern output)
blanchet@39731
   236
  |> Option.map fst
blanchet@39731
   237
blanchet@58359
   238
fun extract_tstplike_proof_and_outcome verbose proof_delims known_failures output =
blanchet@56130
   239
  (case (extract_tstplike_proof proof_delims output,
blanchet@57739
   240
      extract_known_atp_failure known_failures output) of
blanchet@49715
   241
    (_, SOME ProofIncomplete) => ("", NONE)
blanchet@58608
   242
  | (_, SOME ProofUnparsable) => ("", NONE)
blanchet@44087
   243
  | ("", SOME ProofMissing) => ("", NONE)
blanchet@44087
   244
  | ("", NONE) => ("", SOME (UnknownError (short_output verbose output)))
blanchet@49731
   245
  | res as ("", _) => res
blanchet@56130
   246
  | (tstplike_proof, _) => (tstplike_proof, NONE))
blanchet@39692
   247
blanchet@54723
   248
type atp_step_name = string * string list
blanchet@39692
   249
blanchet@43809
   250
fun is_same_atp_step (s1, _) (s2, _) = s1 = s2
blanchet@43809
   251
blanchet@43809
   252
val vampire_fact_prefix = "f"
blanchet@39692
   253
blanchet@53892
   254
fun vampire_step_name_ord p =
blanchet@39695
   255
  let val q = pairself fst p in
blanchet@53892
   256
    (* The "unprefix" part is to cope with Vampire's output. *)
blanchet@56130
   257
    (case pairself (Int.fromString o perhaps (try (unprefix vampire_fact_prefix))) q of
blanchet@53892
   258
      (SOME i, SOME j) => int_ord (i, j)
blanchet@56130
   259
    | _ => raise Fail "not Vampire")
blanchet@39692
   260
  end
blanchet@39692
   261
blanchet@56153
   262
type ('a, 'b) atp_step = atp_step_name * atp_formula_role * 'a * 'b * atp_step_name list
blanchet@39692
   263
blanchet@56153
   264
type 'a atp_proof = (('a, 'a, ('a, 'a atp_type) atp_term, 'a) atp_formula, string) atp_step list
blanchet@39692
   265
blanchet@39692
   266
(**** PARSING OF TSTP FORMAT ****)
blanchet@39692
   267
blanchet@57746
   268
(* Strings enclosed in single quotes (e.g., file names), identifiers possibly starting
blanchet@57746
   269
   with "$" and possibly with "!" in them (for "z3_tptp"). *)
blanchet@39692
   270
val scan_general_id =
blanchet@48932
   271
  $$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode
blanchet@57746
   272
  || (Scan.repeat ($$ "$") -- Scan.many1 Symbol.is_letdig >> (op ^ o pairself implode))
blanchet@57746
   273
    -- Scan.optional (Scan.repeat ($$ "!") -- Scan.many1 Symbol.is_letdig >> (op ^ o pairself implode)) ""
blanchet@57746
   274
    >> op ^
blanchet@39692
   275
blanchet@46104
   276
val skip_term =
blanchet@46079
   277
  let
blanchet@46104
   278
    fun skip _ accum [] = (accum, [])
blanchet@56145
   279
      | skip n accum (ss as s :: ss') =
blanchet@56141
   280
        if s = "," andalso n = 0 then
blanchet@56141
   281
          (accum, ss)
blanchet@57739
   282
        else if member (op =) [")", "]"] s then
blanchet@56145
   283
          if n = 0 then (accum, ss) else skip (n - 1) (s :: accum) ss'
blanchet@57739
   284
        else if member (op =) ["(", "["] s then
blanchet@56145
   285
          skip (n + 1) (s :: accum) ss'
blanchet@56141
   286
        else
blanchet@56145
   287
          skip n (s :: accum) ss'
blanchet@56141
   288
  in
blanchet@56141
   289
    skip 0 [] #>> (rev #> implode)
blanchet@56141
   290
  end
blanchet@46079
   291
blanchet@46079
   292
datatype source =
blanchet@46079
   293
  File_Source of string * string option |
blanchet@46080
   294
  Inference_Source of string * string list
blanchet@46079
   295
blanchet@49147
   296
val dummy_phi = AAtom (ATerm (("", []), []))
blanchet@46104
   297
val dummy_inference = Inference_Source ("", [])
fleury@58597
   298
val dummy_atype = AType(("", []), [])
blanchet@46104
   299
blanchet@56111
   300
(* "skip_term" is there to cope with Waldmeister nonsense such as "theory(equality)". *)
blanchet@51026
   301
fun parse_dependency x =
blanchet@51026
   302
  (parse_inference_source >> snd
blanchet@51026
   303
   || scan_general_id --| skip_term >> single) x
blanchet@51026
   304
and parse_dependencies x =
blanchet@58998
   305
  (Scan.repeat (Scan.option ($$ ",") |-- parse_dependency)
blanchet@58998
   306
   >> (flat #> filter_out (curry (op =) "theory"))) x
blanchet@51026
   307
and parse_file_source x =
blanchet@51026
   308
  (Scan.this_string "file" |-- $$ "(" |-- scan_general_id
fleury@58496
   309
   -- Scan.option ($$ "," |-- scan_general_id
blanchet@58608
   310
     --| Scan.option ($$ "," |-- $$ "[" -- Scan.option scan_general_id --| $$ "]")) --| $$ ")") x
blanchet@51026
   311
and parse_inference_source x =
blanchet@51026
   312
  (Scan.this_string "inference" |-- $$ "(" |-- scan_general_id
blanchet@51026
   313
   --| skip_term --| $$ "," --| skip_term --| $$ "," --| $$ "["
blanchet@51026
   314
   -- parse_dependencies --| $$ "]" --| $$ ")") x
blanchet@52213
   315
and skip_introduced x =
blanchet@54130
   316
  (Scan.this_string "introduced" |-- $$ "(" |-- skip_term
blanchet@54130
   317
   -- Scan.repeat ($$ "," |-- skip_term) --| $$ ")") x
blanchet@51026
   318
and parse_source x =
blanchet@51026
   319
  (parse_file_source >> File_Source
blanchet@51026
   320
   || parse_inference_source >> Inference_Source
blanchet@52213
   321
   || skip_introduced >> K dummy_inference (* for Vampire *)
blanchet@53893
   322
   || scan_general_id >> (fn s => Inference_Source ("", [s])) (* for E *)
blanchet@46104
   323
   || skip_term >> K dummy_inference) x
blanchet@39692
   324
blanchet@56141
   325
fun list_app (f, args) = fold (fn arg => fn f => ATerm ((tptp_app, []), [f, arg])) args f
blanchet@43807
   326
blanchet@56162
   327
fun parse_class x = scan_general_id x
blanchet@56162
   328
and parse_classes x = (parse_class ::: Scan.repeat ($$ "&" |-- parse_class)) x
blanchet@56153
   329
blanchet@56153
   330
fun parse_type x =
blanchet@56162
   331
  (scan_general_id -- Scan.optional ($$ "{" |-- parse_classes --| $$ "}") []
blanchet@56153
   332
    -- Scan.optional ($$ "(" |-- parse_types --| $$ ")") []
blanchet@56153
   333
   >> AType) x
blanchet@56153
   334
and parse_types x = (parse_type ::: Scan.repeat ($$ "," |-- parse_type)) x
blanchet@56153
   335
blanchet@46752
   336
(* We currently ignore TFF and THF types. *)
blanchet@56153
   337
fun parse_type_signature x = Scan.repeat (($$ tptp_has_type || $$ tptp_fun_type) |-- parse_arg) x
blanchet@43809
   338
and parse_arg x =
blanchet@56153
   339
  ($$ "(" |-- parse_term --| $$ ")" --| parse_type_signature
blanchet@56153
   340
   || scan_general_id --| parse_type_signature
blanchet@56153
   341
       -- Scan.optional ($$ "<" |-- parse_types --| $$ ">") []
blanchet@56153
   342
       -- Scan.optional ($$ "(" |-- parse_terms --| $$ ")") []
blanchet@56153
   343
     >> ATerm) x
blanchet@57739
   344
and parse_term x =
blanchet@57739
   345
  (parse_arg -- Scan.repeat ($$ tptp_app |-- parse_arg)
blanchet@57739
   346
   --| Scan.option parse_type_signature >> list_app) x
blanchet@56141
   347
and parse_terms x = (parse_term ::: Scan.repeat ($$ "," |-- parse_term)) x
blanchet@39692
   348
blanchet@39822
   349
fun parse_atom x =
blanchet@56153
   350
  (parse_term -- Scan.option (Scan.option ($$ tptp_not_infix) --| $$ tptp_equal -- parse_term)
blanchet@39822
   351
   >> (fn (u1, NONE) => AAtom u1
blanchet@46752
   352
        | (u1, SOME (neg, u2)) =>
blanchet@49147
   353
          AAtom (ATerm (("equal", []), [u1, u2])) |> is_some neg ? mk_anot)) x
blanchet@39692
   354
blanchet@56130
   355
(* TPTP formulas are fully parenthesized, so we don't need to worry about operator precedence. *)
blanchet@43476
   356
fun parse_literal x =
blanchet@43809
   357
  ((Scan.repeat ($$ tptp_not) >> length)
blanchet@43476
   358
      -- ($$ "(" |-- parse_formula --| $$ ")"
blanchet@43476
   359
          || parse_quantified_formula
blanchet@43476
   360
          || parse_atom)
blanchet@43476
   361
      >> (fn (n, phi) => phi |> n mod 2 = 1 ? mk_anot)) x
blanchet@43476
   362
and parse_formula x =
blanchet@43476
   363
  (parse_literal
blanchet@44004
   364
   -- Scan.option ((Scan.this_string tptp_implies
blanchet@44004
   365
                    || Scan.this_string tptp_iff
blanchet@44004
   366
                    || Scan.this_string tptp_not_iff
blanchet@44004
   367
                    || Scan.this_string tptp_if
blanchet@44004
   368
                    || $$ tptp_or
blanchet@44004
   369
                    || $$ tptp_and) -- parse_formula)
blanchet@39692
   370
   >> (fn (phi1, NONE) => phi1
blanchet@44004
   371
        | (phi1, SOME (c, phi2)) =>
blanchet@44004
   372
          if c = tptp_implies then mk_aconn AImplies phi1 phi2
blanchet@44004
   373
          else if c = tptp_iff then mk_aconn AIff phi1 phi2
blanchet@44004
   374
          else if c = tptp_not_iff then mk_anot (mk_aconn AIff phi1 phi2)
blanchet@44004
   375
          else if c = tptp_if then mk_aconn AImplies phi2 phi1
blanchet@44004
   376
          else if c = tptp_or then mk_aconn AOr phi1 phi2
blanchet@44004
   377
          else if c = tptp_and then mk_aconn AAnd phi1 phi2
blanchet@44004
   378
          else raise Fail ("impossible connective " ^ quote c))) x
blanchet@43476
   379
and parse_quantified_formula x =
blanchet@43809
   380
  (($$ tptp_forall >> K AForall || $$ tptp_exists >> K AExists)
blanchet@43476
   381
   --| $$ "[" -- parse_terms --| $$ "]" --| $$ ":" -- parse_literal
blanchet@56141
   382
   >> (fn ((q, ts), phi) => AQuant (q, map (fn ATerm ((s, []), _) => (s, NONE)) ts, phi))) x
blanchet@39692
   383
blanchet@39692
   384
val parse_tstp_extra_arguments =
blanchet@57739
   385
  Scan.optional ($$ "," |-- parse_source --| Scan.option ($$ "," |-- skip_term)) dummy_inference
blanchet@39692
   386
blanchet@48942
   387
val waldmeister_conjecture_name = "conjecture_1"
blanchet@43784
   388
blanchet@43407
   389
val tofof_fact_prefix = "fof_"
blanchet@41451
   390
blanchet@43784
   391
fun is_same_term subst tm1 tm2 =
blanchet@43784
   392
  let
fleury@58599
   393
    fun do_term_pair (AAbs (((var1, typ1), body1), args1)) (AAbs (((var2, typ2), body2), args2))
fleury@58599
   394
          (SOME subst) =
fleury@58599
   395
        if typ1 <> typ2 andalso length args1 = length args2 then NONE
fleury@58599
   396
        else
fleury@58599
   397
          let val ls = length subst in
fleury@58599
   398
            SOME ((var1, var2) :: subst)
fleury@58599
   399
            |> do_term_pair body1 body2
fleury@58599
   400
            |> (fn SOME subst => SOME (nth_drop (length subst - ls - 1) subst)
fleury@58599
   401
                 | NONE => NONE)
fleury@58599
   402
            |> (if length args1 = length args2
fleury@58599
   403
              then fold2 do_term_pair args1 args2
fleury@58600
   404
              else K NONE)
fleury@58599
   405
          end
fleury@58599
   406
      | do_term_pair (ATerm ((s1, _), args1)) (ATerm ((s2, _), args2)) (SOME subst) =
blanchet@56130
   407
        (case pairself is_tptp_variable (s1, s2) of
blanchet@43784
   408
          (true, true) =>
blanchet@43784
   409
          (case AList.lookup (op =) subst s1 of
fleury@58599
   410
            SOME s2' => if s2' = s2 then SOME subst else NONE
fleury@58599
   411
          | NONE =>
fleury@58599
   412
            if null (AList.find (op =) subst s2) then SOME ((s1, s2) :: subst)
fleury@58599
   413
            else NONE)
blanchet@43784
   414
        | (false, false) =>
fleury@58599
   415
          if s1 = s2 then
fleury@58599
   416
            SOME subst
blanchet@43784
   417
          else
blanchet@43784
   418
            NONE
fleury@58599
   419
        | _ => NONE) |> (if length args1 = length args2
fleury@58599
   420
                       then fold2 do_term_pair args1 args2
fleury@58600
   421
                       else K NONE)
fleury@58599
   422
      | do_term_pair _ _ _ = NONE
blanchet@56130
   423
  in
fleury@58599
   424
    SOME subst |> do_term_pair tm1 tm2 |> is_some
blanchet@56130
   425
  end
blanchet@43784
   426
blanchet@48936
   427
fun is_same_formula comm subst (AQuant (q1, xs1, phi1)) (AQuant (q2, xs2, phi2)) =
blanchet@43784
   428
    q1 = q2 andalso length xs1 = length xs2 andalso
blanchet@48936
   429
    is_same_formula comm ((map fst xs1 ~~ map fst xs2) @ subst) phi1 phi2
blanchet@48936
   430
  | is_same_formula comm subst (AConn (c1, phis1)) (AConn (c2, phis2)) =
blanchet@43784
   431
    c1 = c2 andalso length phis1 = length phis2 andalso
blanchet@48936
   432
    forall (uncurry (is_same_formula comm subst)) (phis1 ~~ phis2)
blanchet@58544
   433
  | is_same_formula comm subst (AAtom (tm1 as ATerm (("equal", tys), [tm11, tm12]))) (AAtom tm2) =
blanchet@48941
   434
    is_same_term subst tm1 tm2 orelse
blanchet@56141
   435
    (comm andalso is_same_term subst (ATerm (("equal", tys), [tm12, tm11])) tm2)
blanchet@48936
   436
  | is_same_formula _ subst (AAtom tm1) (AAtom tm2) = is_same_term subst tm1 tm2
blanchet@48936
   437
  | is_same_formula _ _ _ _ = false
blanchet@43784
   438
blanchet@51536
   439
fun matching_formula_line_identifier phi (Formula ((ident, _), _, phi', _, _)) =
blanchet@48936
   440
    if is_same_formula true [] phi phi' then SOME (ident, phi') else NONE
blanchet@43784
   441
  | matching_formula_line_identifier _ _ = NONE
blanchet@43784
   442
blanchet@56141
   443
fun find_formula_in_problem phi =
blanchet@56141
   444
  maps snd
blanchet@56141
   445
  #> map_filter (matching_formula_line_identifier phi)
blanchet@56141
   446
  #> try (single o hd)
blanchet@56141
   447
  #> the_default []
blanchet@43784
   448
blanchet@56141
   449
fun commute_eq (AAtom (ATerm ((s, tys), tms))) = AAtom (ATerm ((s, tys), rev tms))
blanchet@48961
   450
  | commute_eq _ = raise Fail "expected equation"
blanchet@48936
   451
blanchet@51027
   452
fun role_of_tptp_string "axiom" = Axiom
blanchet@51027
   453
  | role_of_tptp_string "definition" = Definition
blanchet@51027
   454
  | role_of_tptp_string "lemma" = Lemma
blanchet@51027
   455
  | role_of_tptp_string "hypothesis" = Hypothesis
blanchet@51027
   456
  | role_of_tptp_string "conjecture" = Conjecture
blanchet@51027
   457
  | role_of_tptp_string "negated_conjecture" = Negated_Conjecture
blanchet@51027
   458
  | role_of_tptp_string "plain" = Plain
fleury@58597
   459
  | role_of_tptp_string "type" = Type_Role
blanchet@51027
   460
  | role_of_tptp_string _ = Unknown
blanchet@51027
   461
fleury@58598
   462
val tptp_binary_op_list = [tptp_and, tptp_not_and, tptp_or, tptp_not_or, tptp_implies, tptp_if,
fleury@58598
   463
                             tptp_iff, tptp_not_iff, tptp_equal, tptp_app]
fleury@58598
   464
fleury@58597
   465
fun parse_binary_op x =
fleury@58599
   466
  (foldl1 (op ||) (map Scan.this_string tptp_binary_op_list)
fleury@58599
   467
    >> (fn c => if c = tptp_equal then "equal" else c)) x
fleury@58597
   468
fleury@58597
   469
fun parse_thf0_type x =
fleury@58599
   470
  (($$ "(" |-- parse_thf0_type --| $$ ")" || scan_general_id >> (fn t => AType ((t, []), [])))
fleury@58599
   471
     -- Scan.option ($$ tptp_fun_type |-- parse_thf0_type)
fleury@58599
   472
  >> (fn (a, NONE) => a
fleury@58599
   473
       | (a, SOME b) => AFun (a, b))) x
fleury@58597
   474
fleury@58597
   475
fun mk_ho_of_fo_quant q =
fleury@58597
   476
  if q = tptp_forall then tptp_ho_forall
fleury@58597
   477
  else if q = tptp_exists then tptp_ho_exists
fleury@58597
   478
  else raise Fail ("unrecognized quantification: " ^ q)
fleury@58597
   479
fleury@58597
   480
fun remove_thf_app (ATerm ((x, ty), arg)) =
fleury@58597
   481
  if x = tptp_app then
fleury@58597
   482
    (case arg of
fleury@58597
   483
      ATerm ((x, ty), arg) :: t => remove_thf_app (ATerm ((x, ty), map remove_thf_app arg @ t))
fleury@58597
   484
    | [AAbs ((var, tvar), phi), t] =>
fleury@58597
   485
      remove_thf_app (AAbs ((var, tvar), map remove_thf_app phi @ [t])))
fleury@58597
   486
  else ATerm ((x, ty), map remove_thf_app arg)
fleury@58597
   487
  | remove_thf_app (AAbs (((x, ty), arg), t)) = AAbs (((x, ty), remove_thf_app arg), t)
fleury@58597
   488
fleury@58597
   489
fun parse_thf0_typed_var x =
fleury@58597
   490
  (Scan.repeat (scan_general_id -- Scan.option ($$ ":" |-- parse_thf0_type)
fleury@58597
   491
     --| Scan.option (Scan.this_string ","))
fleury@58597
   492
   || $$ "(" |-- parse_thf0_typed_var --| $$ ")") x
fleury@58597
   493
fleury@58599
   494
fun parse_simple_thf0_term x =
fleury@58599
   495
  ((Scan.this_string tptp_forall || Scan.this_string tptp_exists || Scan.this_string tptp_lambda)
fleury@58599
   496
        -- ($$ "[" |-- parse_thf0_typed_var --| $$ "]" --| $$ ":") -- parse_thf0_term
fleury@58599
   497
      >> (fn ((q, ys), t) =>
fleury@58597
   498
          fold_rev
fleury@58597
   499
            (fn (var, ty) => fn r =>
fleury@58599
   500
                AAbs (((var, the_default dummy_atype ty), r), [])
fleury@58599
   501
                |> (if tptp_lambda <> q then
fleury@58599
   502
                      mk_app (q |> mk_ho_of_fo_quant
fleury@58599
   503
                                |> mk_simple_aterm)
fleury@58599
   504
                    else I))
fleury@58599
   505
            ys t)
fleury@58599
   506
  || Scan.this_string tptp_not |-- parse_thf0_term >> mk_app (mk_simple_aterm tptp_not)
fleury@58599
   507
  || scan_general_id --| $$ ":" -- parse_thf0_type >> (fn (var, typ) => ATerm ((var, [typ]), []))
fleury@58599
   508
  || (Scan.this_string tptp_ho_forall || Scan.this_string tptp_ho_exists) >> mk_simple_aterm
fleury@58599
   509
  || $$ "(" |-- parse_thf0_term --| $$ ")"
fleury@58599
   510
  || scan_general_id >> mk_simple_aterm) x
fleury@58599
   511
and parse_thf0_term x =
fleury@58599
   512
  (parse_simple_thf0_term
fleury@58599
   513
      -- Scan.option (parse_binary_op
fleury@58599
   514
      -- parse_thf0_term)
fleury@58599
   515
    >> (fn (t1, SOME (c, t2)) =>
fleury@58599
   516
           if c = tptp_app then mk_app t1 t2 else mk_apps (mk_simple_aterm c) [t1, t2]
fleury@58599
   517
         | (t, NONE) => t)) x
fleury@58597
   518
fleury@58599
   519
fun parse_thf0_formula x = (parse_thf0_term #>> remove_thf_app #>> AAtom) x
fleury@58597
   520
fleury@58599
   521
fun parse_tstp_thf0_line problem =
fleury@58597
   522
  ((Scan.this_string tptp_cnf || Scan.this_string tptp_fof
fleury@58597
   523
    || Scan.this_string tptp_tff || Scan.this_string tptp_thf)
fleury@58597
   524
  -- $$ "(") |-- scan_general_id --| $$ "," -- Symbol.scan_ascii_id --| $$ ","
fleury@58599
   525
  -- parse_thf0_formula -- parse_tstp_extra_arguments --| $$ ")" --| $$ "."
fleury@58597
   526
  >> (fn (((num, role), phi), deps) =>
fleury@58597
   527
      let
fleury@58599
   528
        val role' = role_of_tptp_string role
fleury@58597
   529
        val ((name, phi), rule, deps) =
fleury@58597
   530
          (case deps of
fleury@58599
   531
            File_Source (_, SOME s) =>
fleury@58599
   532
            if role' = Definition then
fleury@58599
   533
              (((num, map fst (find_formula_in_problem phi problem)), phi), "", [])
fleury@58599
   534
            else
fleury@58599
   535
              (((num, [s |> perhaps (try (unprefix tofof_fact_prefix))]), phi), "", [])
fleury@58597
   536
          | Inference_Source (rule, deps) => (((num, []), phi), rule, deps))
fleury@58597
   537
      in
fleury@58599
   538
        [(name, role', phi, rule, map (rpair []) deps)]
fleury@58597
   539
      end)
fleury@58597
   540
blanchet@56130
   541
(* Syntax: (cnf|fof|tff|thf)\(<num>, <formula_role>, <formula> <extra_arguments>\).
blanchet@39692
   542
   The <num> could be an identifier, but we assume integers. *)
blanchet@43784
   543
fun parse_tstp_line problem =
blanchet@43809
   544
  ((Scan.this_string tptp_cnf || Scan.this_string tptp_fof
blanchet@43809
   545
    || Scan.this_string tptp_tff || Scan.this_string tptp_thf) -- $$ "(")
wenzelm@51251
   546
    |-- scan_general_id --| $$ "," -- Symbol.scan_ascii_id --| $$ ","
blanchet@46104
   547
    -- (parse_formula || skip_term >> K dummy_phi) -- parse_tstp_extra_arguments
blanchet@46104
   548
    --| $$ ")" --| $$ "."
blanchet@43784
   549
   >> (fn (((num, role), phi), deps) =>
blanchet@43784
   550
          let
blanchet@48936
   551
            val ((name, phi), rule, deps) =
blanchet@43784
   552
              (* Waldmeister isn't exactly helping. *)
blanchet@56130
   553
              (case deps of
blanchet@46079
   554
                File_Source (_, SOME s) =>
blanchet@48942
   555
                (if s = waldmeister_conjecture_name then
blanchet@56141
   556
                   (case find_formula_in_problem (mk_anot phi) problem of
blanchet@58544
   557
                     (* Waldmeister hack: Get the original orientation of the equation to avoid
blanchet@58544
   558
                        confusing Isar. *)
blanchet@48936
   559
                     [(s, phi')] =>
blanchet@48936
   560
                     ((num, [s]),
blanchet@56130
   561
                      phi |> not (is_same_formula false [] (mk_anot phi) phi') ? commute_eq)
blanchet@56130
   562
                   | _ => ((num, []), phi))
blanchet@48936
   563
                 else
blanchet@48936
   564
                   ((num, [s |> perhaps (try (unprefix tofof_fact_prefix))]),
blanchet@48936
   565
                    phi),
blanchet@48936
   566
                 "", [])
blanchet@46079
   567
              | File_Source _ =>
blanchet@56141
   568
                (((num, map fst (find_formula_in_problem phi problem)), phi), "", [])
blanchet@56130
   569
              | Inference_Source (rule, deps) => (((num, []), phi), rule, deps))
blanchet@56141
   570
            fun mk_step () = (name, role_of_tptp_string role, phi, rule, map (rpair []) deps)
blanchet@43784
   571
          in
blanchet@57746
   572
            [(case role_of_tptp_string role of
blanchet@57746
   573
               Definition =>
blanchet@57746
   574
               (case phi of
fleury@58597
   575
                 AAtom (ATerm (("equal", _), _)) =>
blanchet@57746
   576
                  (* Vampire's equality proxy axiom *)
blanchet@57746
   577
                  (name, Definition, phi, rule, map (rpair []) deps)
fleury@58599
   578
               | _ => mk_step ())
blanchet@57746
   579
             | _ => mk_step ())]
blanchet@43784
   580
          end)
blanchet@39692
   581
blanchet@39692
   582
(**** PARSING OF SPASS OUTPUT ****)
blanchet@39692
   583
blanchet@39692
   584
(* SPASS returns clause references of the form "x.y". We ignore "y", whose role
blanchet@39692
   585
   is not clear anyway. *)
blanchet@39692
   586
val parse_dot_name = scan_general_id --| $$ "." --| scan_general_id
blanchet@39692
   587
blanchet@39692
   588
val parse_spass_annotations =
blanchet@56153
   589
  Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name --| Scan.option ($$ ","))) []
blanchet@39692
   590
blanchet@39692
   591
(* It is not clear why some literals are followed by sequences of stars and/or
blanchet@39692
   592
   pluses. We ignore them. *)
blanchet@39826
   593
fun parse_decorated_atom x =
blanchet@39826
   594
  (parse_atom --| Scan.repeat ($$ "*" || $$ "+" || $$ " ")) x
blanchet@39692
   595
blanchet@49147
   596
fun mk_horn ([], []) = AAtom (ATerm (("c_False", []), []))
blanchet@43784
   597
  | mk_horn ([], pos_lits) = foldr1 (uncurry (mk_aconn AOr)) pos_lits
blanchet@43784
   598
  | mk_horn (neg_lits, []) = mk_anot (foldr1 (uncurry (mk_aconn AAnd)) neg_lits)
blanchet@39692
   599
  | mk_horn (neg_lits, pos_lits) =
blanchet@43784
   600
    mk_aconn AImplies (foldr1 (uncurry (mk_aconn AAnd)) neg_lits)
blanchet@43784
   601
                      (foldr1 (uncurry (mk_aconn AOr)) pos_lits)
blanchet@39692
   602
blanchet@39869
   603
fun parse_horn_clause x =
blanchet@39869
   604
  (Scan.repeat parse_decorated_atom --| $$ "|" --| $$ "|"
blanchet@39869
   605
     -- Scan.repeat parse_decorated_atom --| $$ "-" --| $$ ">"
blanchet@39869
   606
     -- Scan.repeat parse_decorated_atom
blanchet@39869
   607
   >> (mk_horn o apfst (op @))) x
blanchet@39692
   608
blanchet@47218
   609
val parse_spass_debug =
blanchet@56130
   610
  Scan.option ($$ "(" |-- Scan.repeat (scan_general_id --| Scan.option ($$ ",")) --| $$ ")")
blanchet@47218
   611
blanchet@56130
   612
(* Syntax: <num>[0:<inference><annotations>] <atoms> || <atoms> -> <atoms> .
blanchet@47255
   613
           derived from formulae <ident>* *)
blanchet@49020
   614
fun parse_spass_line x =
blanchet@56130
   615
  (parse_spass_debug |-- scan_general_id --| $$ "[" --| Scan.many1 Symbol.is_digit --| $$ ":"
blanchet@56130
   616
     -- Symbol.scan_ascii_id -- parse_spass_annotations --| $$ "]" -- parse_horn_clause --| $$ "."
blanchet@49020
   617
     -- Scan.option (Scan.this_string "derived from formulae "
blanchet@49020
   618
                     |-- Scan.repeat (scan_general_id --| Scan.option ($$ " ")))
blanchet@49020
   619
   >> (fn ((((num, rule), deps), u), names) =>
blanchet@57746
   620
          [((num, these names), Unknown, u, rule, map (rpair []) deps)])) x
blanchet@46033
   621
blanchet@56130
   622
fun parse_spass_pirate_dependency x = (Scan.option ($$ "-") |-- scan_general_id) x
blanchet@56130
   623
fun parse_spass_pirate_dependencies x =
blanchet@56153
   624
  Scan.repeat (parse_spass_pirate_dependency --| Scan.option ($$ "," || $$ " ")) x
blanchet@56130
   625
fun parse_spass_pirate_file_source x =
blanchet@56130
   626
  ((Scan.this_string "Input" || Scan.this_string "Conj") |-- $$ "(" |-- scan_general_id
blanchet@56130
   627
     --| $$ ")") x
blanchet@56130
   628
fun parse_spass_pirate_inference_source x =
blanchet@56178
   629
  (scan_general_id -- ($$ "(" |-- parse_spass_pirate_dependencies --| $$ ")")) x
blanchet@56130
   630
fun parse_spass_pirate_source x =
blanchet@56130
   631
  (parse_spass_pirate_file_source >> (fn s => File_Source ("", SOME s))
blanchet@56130
   632
   || parse_spass_pirate_inference_source >> Inference_Source) x
blanchet@56130
   633
blanchet@56130
   634
(* Syntax: <num> <stuff> || <atoms> -> <atoms> . origin\(<origin>\) *)
blanchet@56130
   635
fun parse_spass_pirate_line x =
blanchet@56130
   636
  (scan_general_id --| Scan.repeat (~$$ "|") -- parse_horn_clause --| $$ "."
blanchet@56130
   637
     --| Scan.this_string "origin" --| $$ "(" -- parse_spass_pirate_source --| $$ ")"
blanchet@56130
   638
   >> (fn ((((num, u), source))) =>
blanchet@56130
   639
     let
blanchet@56130
   640
       val (names, rule, deps) =
blanchet@56130
   641
         (case source of
blanchet@56130
   642
           File_Source (_, SOME s) => ([s], spass_input_rule, [])
blanchet@56130
   643
         | Inference_Source (rule, deps) => ([], rule, deps))
blanchet@56130
   644
     in
blanchet@57746
   645
       [((num, names), Unknown, u, rule, map (rpair []) deps)]
blanchet@56130
   646
     end)) x
blanchet@49554
   647
blanchet@53214
   648
fun core_inference inf fact = ((fact, [fact]), Unknown, dummy_phi, inf, [])
blanchet@53214
   649
blanchet@46033
   650
(* Syntax: <name> *)
blanchet@57746
   651
fun parse_satallax_core_line x =
blanchet@57746
   652
  (scan_general_id --| Scan.option ($$ " ") >> (single o core_inference satallax_core_rule)) x
blanchet@57746
   653
blanchet@57746
   654
(* Syntax: SZS core <name> ... <name> *)
blanchet@57746
   655
fun parse_z3_tptp_core_line x =
blanchet@57746
   656
  (Scan.this_string "SZS core" |-- Scan.repeat ($$ " " |-- scan_general_id)
blanchet@57746
   657
   >> map (core_inference z3_tptp_core_rule)) x
blanchet@39692
   658
fleury@58600
   659
fun parse_line local_name problem =
fleury@58600
   660
  if local_name = leo2N then parse_tstp_thf0_line problem
fleury@58600
   661
  else if local_name = satallaxN then parse_satallax_core_line
fleury@58600
   662
  else if local_name = spassN then parse_spass_line
fleury@58600
   663
  else if local_name = spass_pirateN then parse_spass_pirate_line
fleury@58600
   664
  else if local_name = z3_tptpN then parse_z3_tptp_core_line
fleury@58496
   665
  else parse_tstp_line problem
fleury@58496
   666
fleury@58600
   667
fun parse_proof local_name problem =
blanchet@51605
   668
  strip_spaces_except_between_idents
blanchet@51605
   669
  #> raw_explode
blanchet@51605
   670
  #> Scan.error (!! (fn _ => raise UNRECOGNIZED_ATP_PROOF ())
fleury@58600
   671
       (Scan.finite Symbol.stopper (Scan.repeat1 (parse_line local_name problem) >> flat)))
blanchet@51605
   672
  #> fst
blanchet@39692
   673
blanchet@53214
   674
fun core_of_agsyhol_proof s =
blanchet@56130
   675
  (case split_lines s of
blanchet@53214
   676
    "The transformed problem consists of the following conjectures:" :: conj ::
blanchet@56130
   677
    _ :: proof_term :: _ => SOME (unprefix " " conj :: find_enclosed "<<" ">>" proof_term)
blanchet@56130
   678
  | _ => NONE)
blanchet@53214
   679
blanchet@58998
   680
fun atp_proof_of_tstplike_proof _ _ "" = []
fleury@58600
   681
  | atp_proof_of_tstplike_proof local_prover problem tstp =
blanchet@56130
   682
    (case core_of_agsyhol_proof tstp of
blanchet@56130
   683
      SOME facts => facts |> map (core_inference agsyhol_core_rule)
blanchet@53214
   684
    | NONE =>
blanchet@56153
   685
      tstp ^ "$" (* the $ sign acts as a sentinel (FIXME: needed?) *)
fleury@58600
   686
      |> parse_proof local_prover problem
fleury@58600
   687
      |> local_prover = vampireN ? perhaps (try (sort (vampire_step_name_ord o pairself #1)))
fleury@58600
   688
      |> local_prover = zipperpositionN ? rev)
blanchet@43809
   689
blanchet@43809
   690
fun clean_up_dependencies _ [] = []
blanchet@52338
   691
  | clean_up_dependencies seen ((name, role, u, rule, deps) :: steps) =
blanchet@56153
   692
    (name, role, u, rule, map_filter (fn dep => find_first (is_same_atp_step dep) seen) deps) ::
blanchet@43809
   693
    clean_up_dependencies (name :: seen) steps
blanchet@43809
   694
blanchet@43816
   695
fun clean_up_atp_proof_dependencies proof = clean_up_dependencies [] proof
blanchet@39692
   696
blanchet@51719
   697
fun map_term_names_in_atp_proof f =
blanchet@51719
   698
  let
blanchet@56162
   699
    fun map_type (AType ((s, clss), tys)) = AType ((f s, map f clss), map map_type tys)
blanchet@56153
   700
      | map_type (AFun (ty, ty')) = AFun (map_type ty, map_type ty')
blanchet@56153
   701
      | map_type (APi (ss, ty)) = APi (map f ss, map_type ty)
blanchet@56153
   702
blanchet@56153
   703
    fun map_term (ATerm ((s, tys), ts)) = ATerm ((f s, map map_type tys), map map_term ts)
blanchet@56153
   704
      | map_term (AAbs (((s, ty), tm), args)) =
blanchet@56153
   705
        AAbs (((f s, map_type ty), map_term tm), map map_term args)
blanchet@56153
   706
blanchet@57739
   707
    fun map_formula (AQuant (q, xs, phi)) = AQuant (q, map (apfst f) xs, map_formula phi)
blanchet@56153
   708
      | map_formula (AConn (c, phis)) = AConn (c, map map_formula phis)
blanchet@56153
   709
      | map_formula (AAtom t) = AAtom (map_term t)
blanchet@56153
   710
blanchet@57739
   711
    fun map_step (name, role, phi, rule, deps) = (name, role, map_formula phi, rule, deps)
blanchet@56153
   712
  in
blanchet@56153
   713
    map map_step
blanchet@56153
   714
  end
blanchet@39694
   715
blanchet@58603
   716
fun nasty_name pool s = Symtab.lookup pool s |> the_default s
blanchet@51719
   717
blanchet@39694
   718
fun nasty_atp_proof pool =
blanchet@51719
   719
  not (Symtab.is_empty pool) ? map_term_names_in_atp_proof (nasty_name pool)
blanchet@39694
   720
blanchet@39692
   721
end;