src/HOL/TPTP/atp_problem_import.ML
author blanchet
Fri, 27 Apr 2012 15:24:37 +0200
changeset 47791 c17cc1380642
parent 47788 44b33c1e702e
child 47793 02bdd591eb8f
permissions -rw-r--r--
smaller batches, to play safe
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/TPTP/atp_problem_import.ML
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     3
    Copyright   2012
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     4
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     5
Import TPTP problems as Isabelle terms or goals.
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     6
*)
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     7
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     8
signature ATP_PROBLEM_IMPORT =
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     9
sig
47670
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
    10
  val nitpick_tptp_file : int -> string -> unit
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
    11
  val refute_tptp_file : int -> string -> unit
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
    12
  val sledgehammer_tptp_file : int -> string -> unit
47765
18f37b7aa6a6 more work on CASC setup
blanchet
parents: 47755
diff changeset
    13
  val isabelle_tptp_file : int -> string -> unit
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    14
  val translate_tptp_file : string -> string -> string -> unit
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    15
end;
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    16
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    17
structure ATP_Problem_Import : ATP_PROBLEM_IMPORT =
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    18
struct
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    19
47643
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    20
open ATP_Util
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    21
open ATP_Problem
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    22
open ATP_Proof
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    23
47776
024cf0f7fb6d further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents: 47771
diff changeset
    24
val debug = false
024cf0f7fb6d further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents: 47771
diff changeset
    25
val overlord = false
024cf0f7fb6d further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents: 47771
diff changeset
    26
47643
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    27
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
    28
(** TPTP parsing **)
47643
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    29
47565
762eb0dacaa6 remove old TPTP CNF/FOF parser; always use Nik's new parser
blanchet
parents: 47562
diff changeset
    30
(* cf. "close_form" in "refute.ML" *)
762eb0dacaa6 remove old TPTP CNF/FOF parser; always use Nik's new parser
blanchet
parents: 47562
diff changeset
    31
fun close_form t =
762eb0dacaa6 remove old TPTP CNF/FOF parser; always use Nik's new parser
blanchet
parents: 47562
diff changeset
    32
  fold (fn ((s, i), T) => fn t' =>
762eb0dacaa6 remove old TPTP CNF/FOF parser; always use Nik's new parser
blanchet
parents: 47562
diff changeset
    33
           Logic.all_const T $ Abs (s, T, abstract_over (Var ((s, i), T), t')))
762eb0dacaa6 remove old TPTP CNF/FOF parser; always use Nik's new parser
blanchet
parents: 47562
diff changeset
    34
       (Term.add_vars t []) t
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    35
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
    36
fun read_tptp_file thy postproc file_name =
47644
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
    37
  let
47718
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    38
    fun has_role role (_, role', _, _) = (role' = role)
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
    39
    fun get_prop (_, _, P, _) =
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
    40
      P |> Logic.varify_global |> close_form |> postproc
47644
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
    41
    val path =
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
    42
      Path.explode file_name
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
    43
      |> (fn path =>
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
    44
             path |> not (Path.is_absolute path)
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
    45
                     ? Path.append (Path.explode "$PWD"))
47714
d6683fe037b1 get rid of old parser, hopefully for good
blanchet
parents: 47670
diff changeset
    46
    val ((_, _, problem), thy) =
d6683fe037b1 get rid of old parser, hopefully for good
blanchet
parents: 47670
diff changeset
    47
      TPTP_Interpret.interpret_file true (Path.explode "$TPTP") path [] [] thy
47718
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    48
    val (conjs, defs_and_nondefs) =
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    49
      problem |> List.partition (has_role TPTP_Syntax.Role_Conjecture)
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    50
              ||> List.partition (has_role TPTP_Syntax.Role_Definition)
47644
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
    51
  in
47718
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    52
    (map get_prop conjs, pairself (map get_prop) defs_and_nondefs,
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
    53
     thy |> Theory.checkpoint |> Proof_Context.init_global)
47557
32f35b3d9e42 started integrating Nik's parser into TPTP command-line tools
blanchet
parents: 46325
diff changeset
    54
  end
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    55
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
    56
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    57
(** Nitpick (alias Nitrox) **)
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    58
47718
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    59
fun aptrueprop f ((t0 as @{const Trueprop}) $ t1) = t0 $ f t1
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    60
  | aptrueprop f t = f t
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    61
47670
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
    62
fun nitpick_tptp_file timeout file_name =
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    63
  let
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
    64
    val (conjs, (defs, nondefs), ctxt) = read_tptp_file @{theory} I file_name
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
    65
    val thy = Proof_Context.theory_of ctxt
47718
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    66
    val defs = defs |> map (ATP_Util.extensionalize_term ctxt
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    67
                            #> aptrueprop (open_form I))
47714
d6683fe037b1 get rid of old parser, hopefully for good
blanchet
parents: 47670
diff changeset
    68
    val state = Proof.init ctxt
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    69
    val params =
47557
32f35b3d9e42 started integrating Nik's parser into TPTP command-line tools
blanchet
parents: 46325
diff changeset
    70
      [("card", "1\<emdash>100"),
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    71
       ("box", "false"),
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    72
       ("sat_solver", "smart"),
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    73
       ("max_threads", "1"),
47791
c17cc1380642 smaller batches, to play safe
blanchet
parents: 47788
diff changeset
    74
       ("batch_size", "5"),
47718
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    75
       ("falsify", if null conjs then "false" else "true"),
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    76
       ("verbose", "true"),
47776
024cf0f7fb6d further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents: 47771
diff changeset
    77
       ("debug", if debug then "true" else "false"),
024cf0f7fb6d further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents: 47771
diff changeset
    78
       ("overlord", if overlord then "true" else "false"),
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    79
       ("show_consts", "true"),
47755
df437d1bf8db tweak TPTP Nitpick's output
blanchet
parents: 47718
diff changeset
    80
       ("format", "1"),
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    81
       ("max_potential", "0"),
47718
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    82
       ("timeout", string_of_int timeout),
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    83
       ("tac_timeout", string_of_int ((timeout + 49) div 50))]
47714
d6683fe037b1 get rid of old parser, hopefully for good
blanchet
parents: 47670
diff changeset
    84
      |> Nitpick_Isar.default_params thy
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    85
    val i = 1
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    86
    val n = 1
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    87
    val step = 0
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    88
    val subst = []
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    89
  in
47715
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47714
diff changeset
    90
    Nitpick.pick_nits_in_term state params Nitpick.TPTP i n step subst
47766
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
    91
        defs nondefs (case conjs of conj :: _ => conj | [] => @{prop True});
47718
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
    92
    ()
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    93
  end
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    94
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    95
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    96
(** Refute **)
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    97
47670
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
    98
fun refute_tptp_file timeout file_name =
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    99
  let
47766
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   100
    fun print_szs_from_outcome falsify s =
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   101
      "% SZS status " ^
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   102
      (if s = "genuine" then
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   103
         if falsify then "CounterSatisfiable" else "Satisfiable"
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   104
       else
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   105
         "Unknown")
47788
blanchet
parents: 47785
diff changeset
   106
      |> Output.urgent_message
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   107
    val (conjs, (defs, nondefs), ctxt) = read_tptp_file @{theory} I file_name
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   108
    val params =
47714
d6683fe037b1 get rid of old parser, hopefully for good
blanchet
parents: 47670
diff changeset
   109
      [("maxtime", string_of_int timeout),
d6683fe037b1 get rid of old parser, hopefully for good
blanchet
parents: 47670
diff changeset
   110
       ("maxvars", "100000")]
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   111
  in
47718
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
   112
    Refute.refute_term ctxt params (defs @ nondefs)
47766
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   113
        (case conjs of conj :: _ => conj | [] => @{prop True})
47718
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
   114
    |> print_szs_from_outcome (not (null conjs))
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   115
  end
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   116
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   117
47766
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   118
(** Sledgehammer and Isabelle (combination of provers) **)
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   119
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   120
val cvc3N = "cvc3"
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   121
val z3N = "z3"
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   122
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   123
fun can_tac ctxt tactic conj = can (Goal.prove ctxt [] [] conj) (K tactic)
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   124
47766
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   125
fun SOLVE_TIMEOUT seconds name tac st =
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   126
  let
47788
blanchet
parents: 47785
diff changeset
   127
    val _ = Output.urgent_message ("running " ^ name ^ " for " ^
blanchet
parents: 47785
diff changeset
   128
                                   string_of_int seconds ^ " s")
47766
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   129
    val result =
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   130
      TimeLimit.timeLimit (Time.fromSeconds seconds)
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   131
        (fn () => SINGLE (SOLVE tac) st) ()
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   132
      handle TimeLimit.TimeOut => NONE
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   133
        | ERROR _ => NONE
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   134
  in
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   135
    case result of
47788
blanchet
parents: 47785
diff changeset
   136
      NONE => (Output.urgent_message ("FAILURE: " ^ name); Seq.empty)
blanchet
parents: 47785
diff changeset
   137
    | SOME st' => (Output.urgent_message ("SUCCESS: " ^ name); Seq.single st')
47766
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   138
  end
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   139
47788
blanchet
parents: 47785
diff changeset
   140
fun atp_tac ctxt override_params timeout prover =
47766
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   141
  Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt
47788
blanchet
parents: 47785
diff changeset
   142
      ([("debug", if debug then "true" else "false"),
blanchet
parents: 47785
diff changeset
   143
        ("overlord", if overlord then "true" else "false"),
blanchet
parents: 47785
diff changeset
   144
        ("provers", prover),
blanchet
parents: 47785
diff changeset
   145
        ("timeout", string_of_int timeout)] @ override_params)
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   146
      {add = [], del = [], only = true}
47765
18f37b7aa6a6 more work on CASC setup
blanchet
parents: 47755
diff changeset
   147
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   148
fun sledgehammer_tac ctxt timeout i =
47766
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   149
  let
47788
blanchet
parents: 47785
diff changeset
   150
    fun slice overload_params fraq prover =
blanchet
parents: 47785
diff changeset
   151
      SOLVE_TIMEOUT (timeout div fraq) prover
blanchet
parents: 47785
diff changeset
   152
                    (atp_tac ctxt overload_params timeout prover i)
47766
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   153
  in
47788
blanchet
parents: 47785
diff changeset
   154
    slice [] 7 ATP_Systems.satallaxN
blanchet
parents: 47785
diff changeset
   155
    ORELSE slice [] 7 ATP_Systems.spassN
blanchet
parents: 47785
diff changeset
   156
    ORELSE slice [] 7 z3N
blanchet
parents: 47785
diff changeset
   157
    ORELSE slice [] 7 ATP_Systems.vampireN
blanchet
parents: 47785
diff changeset
   158
    ORELSE slice [] 7 ATP_Systems.eN
blanchet
parents: 47785
diff changeset
   159
    ORELSE slice [] 7 cvc3N
blanchet
parents: 47785
diff changeset
   160
    ORELSE slice [] 7 ATP_Systems.leo2N
47766
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   161
  end
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   162
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   163
fun auto_etc_tac ctxt timeout i =
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   164
  SOLVE_TIMEOUT (timeout div 10) "simp" (asm_full_simp_tac (simpset_of ctxt) i)
47788
blanchet
parents: 47785
diff changeset
   165
  ORELSE SOLVE_TIMEOUT (timeout div 10) "blast" (blast_tac ctxt i)
blanchet
parents: 47785
diff changeset
   166
  ORELSE SOLVE_TIMEOUT (timeout div 20) "auto+spass"
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   167
      (auto_tac ctxt
47788
blanchet
parents: 47785
diff changeset
   168
       THEN ALLGOALS (atp_tac ctxt [] (timeout div 20) ATP_Systems.spassN))
blanchet
parents: 47785
diff changeset
   169
  ORELSE SOLVE_TIMEOUT (timeout div 20) "fast" (fast_tac ctxt i)
blanchet
parents: 47785
diff changeset
   170
  ORELSE SOLVE_TIMEOUT (timeout div 20) "best" (best_tac ctxt i)
blanchet
parents: 47785
diff changeset
   171
  ORELSE SOLVE_TIMEOUT (timeout div 20) "force" (force_tac ctxt i)
blanchet
parents: 47785
diff changeset
   172
  ORELSE SOLVE_TIMEOUT (timeout div 20) "arith" (Arith_Data.arith_tac ctxt i)
blanchet
parents: 47785
diff changeset
   173
  ORELSE SOLVE_TIMEOUT timeout "fastforce" (fast_force_tac ctxt i)
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   174
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   175
val problem_const_prefix = Context.theory_name @{theory} ^ Long_Name.separator
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   176
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   177
(* Isabelle's standard automatic tactics ("auto", etc.) are more eager to
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   178
   unfold "definitions" of free variables than of constants (cf. PUZ107^5). *)
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   179
val freeze_problem_consts =
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   180
  map_aterms (fn t as Const (s, T) =>
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   181
                 if String.isPrefix problem_const_prefix s then
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   182
                   Free (Long_Name.base_name s, T)
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   183
                 else
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   184
                   t
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   185
               | t => t)
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   186
47776
024cf0f7fb6d further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents: 47771
diff changeset
   187
fun make_conj (defs, nondefs) conjs =
024cf0f7fb6d further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents: 47771
diff changeset
   188
  Logic.list_implies (rev defs @ rev nondefs,
024cf0f7fb6d further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents: 47771
diff changeset
   189
                      case conjs of conj :: _ => conj | [] => @{prop False})
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   190
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   191
fun print_szs_from_success conjs success =
47788
blanchet
parents: 47785
diff changeset
   192
  Output.urgent_message ("% SZS status " ^
blanchet
parents: 47785
diff changeset
   193
                         (if success then
blanchet
parents: 47785
diff changeset
   194
                            if null conjs then "Unsatisfiable" else "Theorem"
blanchet
parents: 47785
diff changeset
   195
                          else
blanchet
parents: 47785
diff changeset
   196
                            "% SZS status Unknown"))
47765
18f37b7aa6a6 more work on CASC setup
blanchet
parents: 47755
diff changeset
   197
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   198
fun sledgehammer_tptp_file timeout file_name =
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   199
  let
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   200
    val (conjs, assms, ctxt) =
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   201
      read_tptp_file @{theory} freeze_problem_consts file_name
47776
024cf0f7fb6d further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents: 47771
diff changeset
   202
    val conj = make_conj assms conjs
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   203
  in
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   204
    can_tac ctxt (sledgehammer_tac ctxt timeout 1) conj
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   205
    |> print_szs_from_success conjs
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   206
  end
47766
9f7cdd5fff7c more work on TPTP Isabelle and Sledgehammer tactics
blanchet
parents: 47765
diff changeset
   207
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   208
fun isabelle_tptp_file timeout file_name =
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   209
  let
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   210
    val (conjs, assms, ctxt) =
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   211
      read_tptp_file @{theory} freeze_problem_consts file_name
47776
024cf0f7fb6d further tweaking for Satallax, so that TPTP problems before parsing and after generation are as similar as possible/practical
blanchet
parents: 47771
diff changeset
   212
    val conj = make_conj assms conjs
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   213
  in
47785
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   214
    (can_tac ctxt (sledgehammer_tac ctxt (timeout div 2) 1) conj orelse
d27bb852c430 more tweaking of TPTP/CASC setup
blanchet
parents: 47776
diff changeset
   215
     can_tac ctxt (auto_etc_tac ctxt (timeout div 2) 1) conj orelse
47788
blanchet
parents: 47785
diff changeset
   216
     can_tac ctxt (atp_tac ctxt [] timeout ATP_Systems.satallaxN 1) conj)
47771
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   217
    |> print_szs_from_success conjs
ba321ce6f344 tuning; no need for relevance filter
blanchet
parents: 47766
diff changeset
   218
  end
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   219
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   220
(** Translator between TPTP(-like) file formats **)
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   221
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   222
fun translate_tptp_file format in_file_name out_file_name = ()
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   223
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   224
end;