src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML
author blanchet
Mon, 02 May 2011 14:40:57 +0200
changeset 42613 23b13b1bd565
parent 42609 b5e94b70bc06
child 42614 81953e554197
permissions -rw-r--r--
use strings to encode type systems in ATP module, to reduce the amount of out-of-place information and also to make it easier to print the type system used
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 41066
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_provers.ML
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
     2
    Author:     Fabian Immler, TU Muenchen
32996
d2e48879e65a removed disjunctive group cancellation -- provers run independently;
wenzelm
parents: 32995
diff changeset
     3
    Author:     Makarius
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35867
diff changeset
     4
    Author:     Jasmin Blanchette, TU Muenchen
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
     5
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 41066
diff changeset
     6
Generic prover abstraction for Sledgehammer.
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
     7
*)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
     8
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 41066
diff changeset
     9
signature SLEDGEHAMMER_PROVERS =
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
    10
sig
40181
3788b7adab36 integrated "smt" proof method with Sledgehammer
blanchet
parents: 40132
diff changeset
    11
  type failure = ATP_Proof.failure
38988
483879af0643 finished renaming
blanchet
parents: 38985
diff changeset
    12
  type locality = Sledgehammer_Filter.locality
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    13
  type relevance_fudge = Sledgehammer_Filter.relevance_fudge
40114
blanchet
parents: 40072
diff changeset
    14
  type translated_formula = Sledgehammer_ATP_Translate.translated_formula
41138
eb80538166b6 implemented partially-typed "tags" type encoding
blanchet
parents: 41136
diff changeset
    15
  type type_system = Sledgehammer_ATP_Translate.type_system
40068
ed2869dd9bfa renamed modules
blanchet
parents: 40065
diff changeset
    16
  type minimize_command = Sledgehammer_ATP_Reconstruct.minimize_command
39493
cb2208f2c07d move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents: 39492
diff changeset
    17
42579
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
    18
  datatype rich_type_system =
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
    19
    Dumb_Type_Sys of type_system |
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
    20
    Smart_Type_Sys of bool
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
    21
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35867
diff changeset
    22
  type params =
41208
1b28c43a7074 make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents: 41171
diff changeset
    23
    {debug: bool,
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35867
diff changeset
    24
     verbose: bool,
36143
6490319b1703 added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents: 36064
diff changeset
    25
     overlord: bool,
41208
1b28c43a7074 make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents: 41171
diff changeset
    26
     blocking: bool,
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
    27
     provers: string list,
42579
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
    28
     type_sys: rich_type_system,
42180
a6c141925a8a added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents: 42100
diff changeset
    29
     relevance_thresholds: real * real,
a6c141925a8a added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents: 42100
diff changeset
    30
     max_relevant: int option,
a6c141925a8a added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents: 42100
diff changeset
    31
     monomorphize_limit: int,
36235
61159615a0c5 added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents: 36231
diff changeset
    32
     explicit_apply: bool,
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35867
diff changeset
    33
     isar_proof: bool,
36924
ff01d3ae9ad4 renamed options
blanchet
parents: 36922
diff changeset
    34
     isar_shrink_factor: int,
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
    35
     slicing: bool,
38985
162bbbea4e4d added "expect" feature of Nitpick to Sledgehammer, for regression testing
blanchet
parents: 38982
diff changeset
    36
     timeout: Time.time,
162bbbea4e4d added "expect" feature of Nitpick to Sledgehammer, for regression testing
blanchet
parents: 38982
diff changeset
    37
     expect: string}
39493
cb2208f2c07d move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents: 39492
diff changeset
    38
41090
b98fe4de1ecd renamings
blanchet
parents: 41089
diff changeset
    39
  datatype prover_fact =
b98fe4de1ecd renamings
blanchet
parents: 41089
diff changeset
    40
    Untranslated_Fact of (string * locality) * thm |
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
    41
    SMT_Weighted_Fact of (string * locality) * (int option * thm)
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    42
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    43
  type prover_problem =
39318
ad9a1f9b0558 implemented Auto Sledgehammer
blanchet
parents: 39263
diff changeset
    44
    {state: Proof.state,
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38997
diff changeset
    45
     goal: thm,
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38997
diff changeset
    46
     subgoal: int,
40065
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40064
diff changeset
    47
     subgoal_count: int,
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
    48
     facts: prover_fact list,
41741
839d1488045f renamed field
blanchet
parents: 41727
diff changeset
    49
     smt_filter: (string * locality) SMT_Solver.smt_filter_data option}
39493
cb2208f2c07d move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents: 39492
diff changeset
    50
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    51
  type prover_result =
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
    52
    {outcome: failure option,
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40202
diff changeset
    53
     used_facts: (string * locality) list,
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    54
     run_time_in_msecs: int option,
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    55
     message: string}
39493
cb2208f2c07d move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents: 39492
diff changeset
    56
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    57
  type prover = params -> minimize_command -> prover_problem -> prover_result
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
    58
41168
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
    59
  (* for experimentation purposes -- do not use in production code *)
41335
66edbd0f7a2e added "smt_triggers" option to experiment with triggers in Sledgehammer;
blanchet
parents: 41318
diff changeset
    60
  val smt_triggers : bool Unsynchronized.ref
41256
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
    61
  val smt_weights : bool Unsynchronized.ref
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
    62
  val smt_weight_min_facts : int Unsynchronized.ref
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
    63
  val smt_min_weight : int Unsynchronized.ref
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
    64
  val smt_max_weight : int Unsynchronized.ref
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
    65
  val smt_max_weight_index : int Unsynchronized.ref
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
    66
  val smt_weight_curve : (int -> int) Unsynchronized.ref
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
    67
  val smt_max_slices : int Unsynchronized.ref
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
    68
  val smt_slice_fact_frac : real Unsynchronized.ref
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
    69
  val smt_slice_time_frac : real Unsynchronized.ref
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
    70
  val smt_slice_min_secs : int Unsynchronized.ref
41168
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
    71
41089
2e69fb6331cb moved function to later module
blanchet
parents: 41088
diff changeset
    72
  val das_Tool : string
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
    73
  val select_smt_solver : string -> Proof.context -> Proof.context
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 41066
diff changeset
    74
  val is_smt_prover : Proof.context -> string -> bool
41727
ab3f6d76fb23 available_provers ~> supported_provers (for clarity)
blanchet
parents: 41723
diff changeset
    75
  val is_prover_supported : Proof.context -> string -> bool
40072
27f2a45b0aab more robust handling of "remote_" vs. non-"remote_" provers
blanchet
parents: 40071
diff changeset
    76
  val is_prover_installed : Proof.context -> string -> bool
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
    77
  val default_max_relevant_for_prover : Proof.context -> bool -> string -> int
40369
53dca3bd4250 use the SMT integration's official list of built-ins
blanchet
parents: 40341
diff changeset
    78
  val is_built_in_const_for_prover :
41336
0ea5b9c7d233 proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents: 41335
diff changeset
    79
    Proof.context -> string -> string * typ -> term list -> bool * term list
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 41066
diff changeset
    80
  val atp_relevance_fudge : relevance_fudge
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 41066
diff changeset
    81
  val smt_relevance_fudge : relevance_fudge
40941
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40723
diff changeset
    82
  val relevance_fudge_for_prover : Proof.context -> string -> relevance_fudge
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
    83
  val dest_dir : string Config.T
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
    84
  val problem_prefix : string Config.T
39003
c2aebd79981f run relevance filter in a thread, to avoid blocking
blanchet
parents: 39000
diff changeset
    85
  val measure_run_time : bool Config.T
41256
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
    86
  val weight_smt_fact :
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
    87
    theory -> int -> ((string * locality) * thm) * int
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
    88
    -> (string * locality) * (int option * thm)
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: 42588
diff changeset
    89
  val is_rich_type_sys_fairly_sound : rich_type_system -> bool
41091
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
    90
  val untranslated_fact : prover_fact -> (string * locality) * thm
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
    91
  val smt_weighted_fact :
41256
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
    92
    theory -> int -> prover_fact * int
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
    93
    -> (string * locality) * (int option * thm)
41727
ab3f6d76fb23 available_provers ~> supported_provers (for clarity)
blanchet
parents: 41723
diff changeset
    94
  val supported_provers : Proof.context -> unit
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
    95
  val kill_provers : unit -> unit
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
    96
  val running_provers : unit -> unit
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
    97
  val messages : int option -> unit
42444
8e5438dc70bb cleanup: get rid of "may_slice" arguments without changing semantics
blanchet
parents: 42443
diff changeset
    98
  val get_prover : Proof.context -> bool -> string -> prover
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
    99
  val setup : theory -> theory
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
   100
end;
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
   101
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 41066
diff changeset
   102
structure Sledgehammer_Provers : SLEDGEHAMMER_PROVERS =
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
   103
struct
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
   104
38028
22dcaec5fa77 minor refactoring
blanchet
parents: 38023
diff changeset
   105
open ATP_Problem
39491
2416666e6f94 refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents: 39453
diff changeset
   106
open ATP_Proof
38028
22dcaec5fa77 minor refactoring
blanchet
parents: 38023
diff changeset
   107
open ATP_Systems
39494
bf7dd4902321 rename "Metis_Clauses" to "Metis_Translate" for consistency with "Sledgehammer_Translate"
blanchet
parents: 39493
diff changeset
   108
open Metis_Translate
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   109
open Sledgehammer_Util
38988
483879af0643 finished renaming
blanchet
parents: 38985
diff changeset
   110
open Sledgehammer_Filter
40068
ed2869dd9bfa renamed modules
blanchet
parents: 40065
diff changeset
   111
open Sledgehammer_ATP_Translate
ed2869dd9bfa renamed modules
blanchet
parents: 40065
diff changeset
   112
open Sledgehammer_ATP_Reconstruct
37583
9ce2451647d5 factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents: 37581
diff changeset
   113
9ce2451647d5 factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents: 37581
diff changeset
   114
(** The Sledgehammer **)
9ce2451647d5 factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents: 37581
diff changeset
   115
38102
019a49759829 fix bug in the newly introduced "bound concealing" code
blanchet
parents: 38100
diff changeset
   116
(* Identifier to distinguish Sledgehammer from other tools using
019a49759829 fix bug in the newly introduced "bound concealing" code
blanchet
parents: 38100
diff changeset
   117
   "Async_Manager". *)
37585
c2ed8112ce57 multiplexing
blanchet
parents: 37584
diff changeset
   118
val das_Tool = "Sledgehammer"
c2ed8112ce57 multiplexing
blanchet
parents: 37584
diff changeset
   119
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   120
val select_smt_solver =
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41336
diff changeset
   121
  Context.proof_map o SMT_Config.select_solver
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   122
40941
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40723
diff changeset
   123
fun is_smt_prover ctxt name =
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41336
diff changeset
   124
  member (op =) (SMT_Solver.available_solvers_of ctxt) name
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   125
41727
ab3f6d76fb23 available_provers ~> supported_provers (for clarity)
blanchet
parents: 41723
diff changeset
   126
fun is_prover_supported ctxt name =
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42193
diff changeset
   127
  let val thy = Proof_Context.theory_of ctxt in
41727
ab3f6d76fb23 available_provers ~> supported_provers (for clarity)
blanchet
parents: 41723
diff changeset
   128
    is_smt_prover ctxt name orelse member (op =) (supported_atps thy) name
40941
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40723
diff changeset
   129
  end
40072
27f2a45b0aab more robust handling of "remote_" vs. non-"remote_" provers
blanchet
parents: 40071
diff changeset
   130
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41336
diff changeset
   131
fun is_prover_installed ctxt =
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42193
diff changeset
   132
  is_smt_prover ctxt orf is_atp_installed (Proof_Context.theory_of ctxt)
40941
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40723
diff changeset
   133
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   134
fun get_slices slicing slices =
42447
blanchet
parents: 42445
diff changeset
   135
  (0 upto length slices - 1) ~~ slices |> not slicing ? (List.last #> single)
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   136
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   137
fun default_max_relevant_for_prover ctxt slicing name =
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42193
diff changeset
   138
  let val thy = Proof_Context.theory_of ctxt in
40982
d06ffd777f1f honor the default max relevant facts setting from the SMT solvers in Sledgehammer
blanchet
parents: 40979
diff changeset
   139
    if is_smt_prover ctxt name then
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41336
diff changeset
   140
      SMT_Solver.default_max_relevant ctxt name
40982
d06ffd777f1f honor the default max relevant facts setting from the SMT solvers in Sledgehammer
blanchet
parents: 40979
diff changeset
   141
    else
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   142
      fold (Integer.max o snd o snd o snd)
42578
1eaf4d437d4c define type system in ATP module so that ATPs can suggest a type system
blanchet
parents: 42577
diff changeset
   143
           (get_slices slicing (#best_slices (get_atp thy name) ())) 0
40941
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40723
diff changeset
   144
  end
40063
d086e3699e78 bring ATPs and SMT solvers more in line with each other
blanchet
parents: 40062
diff changeset
   145
41140
9c68004b8c9d added Sledgehammer support for higher-order propositional reasoning
blanchet
parents: 41138
diff changeset
   146
(* These are either simplified away by "Meson.presimplify" (most of the time) or
9c68004b8c9d added Sledgehammer support for higher-order propositional reasoning
blanchet
parents: 41138
diff changeset
   147
   handled specially via "fFalse", "fTrue", ..., "fequal". *)
40071
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   148
val atp_irrelevant_consts =
41140
9c68004b8c9d added Sledgehammer support for higher-order propositional reasoning
blanchet
parents: 41138
diff changeset
   149
  [@{const_name False}, @{const_name True}, @{const_name Not},
9c68004b8c9d added Sledgehammer support for higher-order propositional reasoning
blanchet
parents: 41138
diff changeset
   150
   @{const_name conj}, @{const_name disj}, @{const_name implies},
9c68004b8c9d added Sledgehammer support for higher-order propositional reasoning
blanchet
parents: 41138
diff changeset
   151
   @{const_name HOL.eq}, @{const_name If}, @{const_name Let}]
40206
8819ffd60357 better list of irrelevant SMT constants
blanchet
parents: 40205
diff changeset
   152
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   153
fun is_built_in_const_for_prover ctxt name =
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   154
  if is_smt_prover ctxt name then
41336
0ea5b9c7d233 proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents: 41335
diff changeset
   155
    let val ctxt = ctxt |> select_smt_solver name in
0ea5b9c7d233 proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents: 41335
diff changeset
   156
      fn x => fn ts =>
0ea5b9c7d233 proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents: 41335
diff changeset
   157
         if SMT_Builtin.is_builtin_num_ext ctxt (list_comb (Const x, ts)) then
0ea5b9c7d233 proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents: 41335
diff changeset
   158
           (true, [])
0ea5b9c7d233 proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents: 41335
diff changeset
   159
         else if SMT_Builtin.is_builtin_fun_ext ctxt x ts then
0ea5b9c7d233 proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents: 41335
diff changeset
   160
           (true, ts)
0ea5b9c7d233 proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents: 41335
diff changeset
   161
         else
0ea5b9c7d233 proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents: 41335
diff changeset
   162
           (false, ts)
0ea5b9c7d233 proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents: 41335
diff changeset
   163
    end
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   164
  else
41336
0ea5b9c7d233 proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents: 41335
diff changeset
   165
    fn (s, _) => fn ts => (member (op =) atp_irrelevant_consts s, ts)
40071
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   166
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   167
(* FUDGE *)
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   168
val atp_relevance_fudge =
41159
1e12d6495423 honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents: 41152
diff changeset
   169
  {allow_ext = true,
41790
56dcd46ddf7a give more weight to Frees than to Consts in relevance filter
blanchet
parents: 41744
diff changeset
   170
   local_const_multiplier = 1.5,
41159
1e12d6495423 honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents: 41152
diff changeset
   171
   worse_irrel_freq = 100.0,
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   172
   higher_order_irrel_weight = 1.05,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   173
   abs_rel_weight = 0.5,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   174
   abs_irrel_weight = 2.0,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   175
   skolem_irrel_weight = 0.75,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   176
   theory_const_rel_weight = 0.5,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   177
   theory_const_irrel_weight = 0.25,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   178
   intro_bonus = 0.15,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   179
   elim_bonus = 0.15,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   180
   simp_bonus = 0.15,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   181
   local_bonus = 0.55,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   182
   assum_bonus = 1.05,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   183
   chained_bonus = 1.5,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   184
   max_imperfect = 11.5,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   185
   max_imperfect_exp = 1.0,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   186
   threshold_divisor = 2.0,
41093
dfbc8759415f lower fudge factor
blanchet
parents: 41091
diff changeset
   187
   ridiculous_threshold = 0.01}
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   188
40071
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   189
(* FUDGE (FIXME) *)
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   190
val smt_relevance_fudge =
41159
1e12d6495423 honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents: 41152
diff changeset
   191
  {allow_ext = false,
41790
56dcd46ddf7a give more weight to Frees than to Consts in relevance filter
blanchet
parents: 41744
diff changeset
   192
   local_const_multiplier = #local_const_multiplier atp_relevance_fudge,
41159
1e12d6495423 honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents: 41152
diff changeset
   193
   worse_irrel_freq = #worse_irrel_freq atp_relevance_fudge,
40071
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   194
   higher_order_irrel_weight = #higher_order_irrel_weight atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   195
   abs_rel_weight = #abs_rel_weight atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   196
   abs_irrel_weight = #abs_irrel_weight atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   197
   skolem_irrel_weight = #skolem_irrel_weight atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   198
   theory_const_rel_weight = #theory_const_rel_weight atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   199
   theory_const_irrel_weight = #theory_const_irrel_weight atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   200
   intro_bonus = #intro_bonus atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   201
   elim_bonus = #elim_bonus atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   202
   simp_bonus = #simp_bonus atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   203
   local_bonus = #local_bonus atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   204
   assum_bonus = #assum_bonus atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   205
   chained_bonus = #chained_bonus atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   206
   max_imperfect = #max_imperfect atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   207
   max_imperfect_exp = #max_imperfect_exp atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   208
   threshold_divisor = #threshold_divisor atp_relevance_fudge,
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   209
   ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge}
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   210
40941
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40723
diff changeset
   211
fun relevance_fudge_for_prover ctxt name =
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40723
diff changeset
   212
  if is_smt_prover ctxt name then smt_relevance_fudge else atp_relevance_fudge
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   213
41727
ab3f6d76fb23 available_provers ~> supported_provers (for clarity)
blanchet
parents: 41723
diff changeset
   214
fun supported_provers ctxt =
40060
5ef6747aa619 first step in adding support for an SMT backend to Sledgehammer
blanchet
parents: 40059
diff changeset
   215
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42193
diff changeset
   216
    val thy = Proof_Context.theory_of ctxt
40060
5ef6747aa619 first step in adding support for an SMT backend to Sledgehammer
blanchet
parents: 40059
diff changeset
   217
    val (remote_provers, local_provers) =
41727
ab3f6d76fb23 available_provers ~> supported_provers (for clarity)
blanchet
parents: 41723
diff changeset
   218
      sort_strings (supported_atps thy) @
ab3f6d76fb23 available_provers ~> supported_provers (for clarity)
blanchet
parents: 41723
diff changeset
   219
      sort_strings (SMT_Solver.available_solvers_of ctxt)
40060
5ef6747aa619 first step in adding support for an SMT backend to Sledgehammer
blanchet
parents: 40059
diff changeset
   220
      |> List.partition (String.isPrefix remote_prefix)
5ef6747aa619 first step in adding support for an SMT backend to Sledgehammer
blanchet
parents: 40059
diff changeset
   221
  in
41727
ab3f6d76fb23 available_provers ~> supported_provers (for clarity)
blanchet
parents: 41723
diff changeset
   222
    Output.urgent_message ("Supported provers: " ^
40205
277508b07418 if "debug" is on, print list of relevant facts (poweruser request);
blanchet
parents: 40204
diff changeset
   223
                           commas (local_provers @ remote_provers) ^ ".")
40060
5ef6747aa619 first step in adding support for an SMT backend to Sledgehammer
blanchet
parents: 40059
diff changeset
   224
  end
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35867
diff changeset
   225
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
   226
fun kill_provers () = Async_Manager.kill_threads das_Tool "provers"
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
   227
fun running_provers () = Async_Manager.running_threads das_Tool "provers"
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
   228
val messages = Async_Manager.thread_messages das_Tool "prover"
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
   229
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
   230
(** problems, results, ATPs, etc. **)
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35867
diff changeset
   231
42579
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   232
datatype rich_type_system =
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   233
  Dumb_Type_Sys of type_system |
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   234
  Smart_Type_Sys of bool
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   235
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35867
diff changeset
   236
type params =
41208
1b28c43a7074 make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents: 41171
diff changeset
   237
  {debug: bool,
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35867
diff changeset
   238
   verbose: bool,
36143
6490319b1703 added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents: 36064
diff changeset
   239
   overlord: bool,
41208
1b28c43a7074 make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents: 41171
diff changeset
   240
   blocking: bool,
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
   241
   provers: string list,
42579
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   242
   type_sys: rich_type_system,
42180
a6c141925a8a added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents: 42100
diff changeset
   243
   relevance_thresholds: real * real,
a6c141925a8a added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents: 42100
diff changeset
   244
   max_relevant: int option,
a6c141925a8a added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents: 42100
diff changeset
   245
   monomorphize_limit: int,
36235
61159615a0c5 added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents: 36231
diff changeset
   246
   explicit_apply: bool,
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35867
diff changeset
   247
   isar_proof: bool,
36924
ff01d3ae9ad4 renamed options
blanchet
parents: 36922
diff changeset
   248
   isar_shrink_factor: int,
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   249
   slicing: bool,
38985
162bbbea4e4d added "expect" feature of Nitpick to Sledgehammer, for regression testing
blanchet
parents: 38982
diff changeset
   250
   timeout: Time.time,
162bbbea4e4d added "expect" feature of Nitpick to Sledgehammer, for regression testing
blanchet
parents: 38982
diff changeset
   251
   expect: string}
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
   252
41090
b98fe4de1ecd renamings
blanchet
parents: 41089
diff changeset
   253
datatype prover_fact =
b98fe4de1ecd renamings
blanchet
parents: 41089
diff changeset
   254
  Untranslated_Fact of (string * locality) * thm |
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   255
  SMT_Weighted_Fact of (string * locality) * (int option * thm)
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   256
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   257
type prover_problem =
39318
ad9a1f9b0558 implemented Auto Sledgehammer
blanchet
parents: 39263
diff changeset
   258
  {state: Proof.state,
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38997
diff changeset
   259
   goal: thm,
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38997
diff changeset
   260
   subgoal: int,
40065
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40064
diff changeset
   261
   subgoal_count: int,
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   262
   facts: prover_fact list,
41741
839d1488045f renamed field
blanchet
parents: 41727
diff changeset
   263
   smt_filter: (string * locality) SMT_Solver.smt_filter_data option}
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
   264
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   265
type prover_result =
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
   266
  {outcome: failure option,
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35867
diff changeset
   267
   message: string,
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40202
diff changeset
   268
   used_facts: (string * locality) list,
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   269
   run_time_in_msecs: int option}
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
   270
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   271
type prover = params -> minimize_command -> prover_problem -> prover_result
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
   272
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   273
(* configuration attributes *)
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   274
38991
0e2798f30087 rename sledgehammer config attributes
blanchet
parents: 38988
diff changeset
   275
val (dest_dir, dest_dir_setup) =
39003
c2aebd79981f run relevance filter in a thread, to avoid blocking
blanchet
parents: 39000
diff changeset
   276
  Attrib.config_string "sledgehammer_dest_dir" (K "")
38991
0e2798f30087 rename sledgehammer config attributes
blanchet
parents: 38988
diff changeset
   277
  (* Empty string means create files in Isabelle's temporary files directory. *)
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   278
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   279
val (problem_prefix, problem_prefix_setup) =
39003
c2aebd79981f run relevance filter in a thread, to avoid blocking
blanchet
parents: 39000
diff changeset
   280
  Attrib.config_string "sledgehammer_problem_prefix" (K "prob")
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   281
39003
c2aebd79981f run relevance filter in a thread, to avoid blocking
blanchet
parents: 39000
diff changeset
   282
val (measure_run_time, measure_run_time_setup) =
c2aebd79981f run relevance filter in a thread, to avoid blocking
blanchet
parents: 39000
diff changeset
   283
  Attrib.config_bool "sledgehammer_measure_run_time" (K false)
28484
4ed9239b09c1 misc simplifcation and tuning;
wenzelm
parents: 28478
diff changeset
   284
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   285
fun with_path cleanup after f path =
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   286
  Exn.capture f path
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   287
  |> tap (fn _ => cleanup path)
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   288
  |> Exn.release
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   289
  |> tap (after path)
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   290
40060
5ef6747aa619 first step in adding support for an SMT backend to Sledgehammer
blanchet
parents: 40059
diff changeset
   291
fun proof_banner auto =
40224
7883fefd6a7b clear identification;
blanchet
parents: 40207
diff changeset
   292
  if auto then "Auto Sledgehammer found a proof" else "Try this command"
40060
5ef6747aa619 first step in adding support for an SMT backend to Sledgehammer
blanchet
parents: 40059
diff changeset
   293
41723
bb366da22483 enable SMT weights and triggers, since they lead to slight improvements according to the Judgment Day suite
blanchet
parents: 41432
diff changeset
   294
val smt_triggers = Unsynchronized.ref true
bb366da22483 enable SMT weights and triggers, since they lead to slight improvements according to the Judgment Day suite
blanchet
parents: 41432
diff changeset
   295
val smt_weights = Unsynchronized.ref true
41256
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   296
val smt_weight_min_facts = Unsynchronized.ref 20
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   297
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   298
(* FUDGE *)
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   299
val smt_min_weight = Unsynchronized.ref 0
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   300
val smt_max_weight = Unsynchronized.ref 10
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   301
val smt_max_weight_index = Unsynchronized.ref 200
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   302
val smt_weight_curve = Unsynchronized.ref (fn x : int => x * x)
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   303
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   304
fun smt_fact_weight j num_facts =
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   305
  if !smt_weights andalso num_facts >= !smt_weight_min_facts then
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   306
    SOME (!smt_max_weight
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   307
          - (!smt_max_weight - !smt_min_weight + 1)
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   308
            * !smt_weight_curve (Int.max (0, !smt_max_weight_index - j - 1))
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   309
            div !smt_weight_curve (!smt_max_weight_index))
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   310
  else
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   311
    NONE
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   312
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   313
fun weight_smt_fact thy num_facts ((info, th), j) =
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   314
  (info, (smt_fact_weight j num_facts, th |> Thm.transfer thy))
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   315
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: 42588
diff changeset
   316
fun is_rich_type_sys_fairly_sound (Dumb_Type_Sys type_sys) =
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: 42588
diff changeset
   317
    is_type_sys_fairly_sound type_sys
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: 42588
diff changeset
   318
  | is_rich_type_sys_fairly_sound (Smart_Type_Sys full_types) = full_types
42579
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   319
41091
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
   320
fun untranslated_fact (Untranslated_Fact p) = p
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   321
  | untranslated_fact (SMT_Weighted_Fact (info, (_, th))) = (info, th)
42579
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   322
fun atp_translated_fact ctxt fact =
42544
75cb06eee990 reimplemented the hAPP introduction code so that it's done earlier, when the types are still available
blanchet
parents: 42542
diff changeset
   323
    translate_atp_fact ctxt false (untranslated_fact fact)
41256
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   324
fun smt_weighted_fact _ _ (SMT_Weighted_Fact p, _) = p
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   325
  | smt_weighted_fact thy num_facts (fact, j) =
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   326
    (untranslated_fact fact, j) |> weight_smt_fact thy num_facts
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   327
41313
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   328
fun overlord_file_location_for_prover prover =
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   329
  (getenv "ISABELLE_HOME_USER", "prob_" ^ prover)
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   330
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   331
41256
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   332
(* generic TPTP-based ATPs *)
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   333
40409
3642dc3b72e8 invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents: 40370
diff changeset
   334
fun int_opt_add (SOME m) (SOME n) = SOME (m + n)
3642dc3b72e8 invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents: 40370
diff changeset
   335
  | int_opt_add _ _ = NONE
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   336
42452
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   337
val atp_blacklist_max_iters = 10
39492
b1172d65dd28 skip some "important" messages
blanchet
parents: 39491
diff changeset
   338
(* Important messages are important but not so important that users want to see
b1172d65dd28 skip some "important" messages
blanchet
parents: 39491
diff changeset
   339
   them each time. *)
42609
b5e94b70bc06 fixed random number invocation
blanchet
parents: 42593
diff changeset
   340
val atp_important_message_keep_quotient = 10
39492
b1172d65dd28 skip some "important" messages
blanchet
parents: 39491
diff changeset
   341
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: 42588
diff changeset
   342
val fallback_best_type_systems =
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: 42588
diff changeset
   343
  [Preds (Polymorphic, Const_Arg_Types), Many_Typed]
42579
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   344
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   345
fun adjust_dumb_type_sys formats Many_Typed =
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   346
    if member (op =) formats Tff then (Tff, Many_Typed)
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: 42588
diff changeset
   347
    else (Fof, Preds (Mangled_Monomorphic, All_Types))
42579
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   348
  | adjust_dumb_type_sys formats type_sys =
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   349
    if member (op =) formats Fof then (Fof, type_sys)
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   350
    else (Tff, Many_Typed)
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   351
fun determine_format_and_type_sys _ formats (Dumb_Type_Sys type_sys) =
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   352
    adjust_dumb_type_sys formats type_sys
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   353
  | determine_format_and_type_sys best_type_systems formats
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   354
                                  (Smart_Type_Sys full_types) =
42613
23b13b1bd565 use strings to encode type systems in ATP module, to reduce the amount of out-of-place information and also to make it easier to print the type system used
blanchet
parents: 42609
diff changeset
   355
    map type_sys_from_string best_type_systems @ fallback_best_type_systems
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: 42588
diff changeset
   356
    |> find_first (if full_types then is_type_sys_virtually_sound else K true)
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: 42588
diff changeset
   357
    |> the |> adjust_dumb_type_sys formats
42548
ea2a28b1938f make sure the minimizer monomorphizes when it should
blanchet
parents: 42544
diff changeset
   358
42444
8e5438dc70bb cleanup: get rid of "may_slice" arguments without changing semantics
blanchet
parents: 42443
diff changeset
   359
fun run_atp auto name
42578
1eaf4d437d4c define type system in ATP module so that ATPs can suggest a type system
blanchet
parents: 42577
diff changeset
   360
        ({exec, required_execs, arguments, proof_delims, known_failures,
42579
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   361
          hypothesis_kind, formats, best_slices, best_type_systems, ...}
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   362
         : atp_config)
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   363
        ({debug, verbose, overlord, type_sys, max_relevant, monomorphize_limit,
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   364
          explicit_apply, isar_proof, isar_shrink_factor, slicing, timeout, ...}
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   365
         : params)
40983
07526f546680 handle "max_relevant" uniformly
blanchet
parents: 40982
diff changeset
   366
        minimize_command ({state, goal, subgoal, facts, ...} : prover_problem) =
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   367
  let
42182
a630978fc967 start monomorphization process with subgoal, not entire goal, to avoid needless instances (and only print monomorphization messages in debug mode)
blanchet
parents: 42181
diff changeset
   368
    val thy = Proof.theory_of state
39318
ad9a1f9b0558 implemented Auto Sledgehammer
blanchet
parents: 39263
diff changeset
   369
    val ctxt = Proof.context_of state
42579
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   370
    val (format, type_sys) =
2552c09b1a72 implement the new ATP type system in Sledgehammer
blanchet
parents: 42578
diff changeset
   371
      determine_format_and_type_sys best_type_systems formats type_sys
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38997
diff changeset
   372
    val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal
41159
1e12d6495423 honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents: 41152
diff changeset
   373
    val (dest_dir, problem_prefix) =
1e12d6495423 honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents: 41152
diff changeset
   374
      if overlord then overlord_file_location_for_prover name
1e12d6495423 honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents: 41152
diff changeset
   375
      else (Config.get ctxt dest_dir, Config.get ctxt problem_prefix)
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   376
    val problem_file_name =
41159
1e12d6495423 honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents: 41152
diff changeset
   377
      Path.basic (problem_prefix ^ (if overlord then "" else serial_string ()) ^
1e12d6495423 honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents: 41152
diff changeset
   378
                  "_" ^ string_of_int subgoal)
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   379
    val problem_path_name =
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
   380
      if dest_dir = "" then
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   381
        File.tmp_path problem_file_name
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
   382
      else if File.exists (Path.explode dest_dir) then
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   383
        Path.append (Path.explode dest_dir) problem_file_name
39003
c2aebd79981f run relevance filter in a thread, to avoid blocking
blanchet
parents: 39000
diff changeset
   384
      else
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
   385
        error ("No such directory: " ^ quote dest_dir ^ ".")
39003
c2aebd79981f run relevance filter in a thread, to avoid blocking
blanchet
parents: 39000
diff changeset
   386
    val measure_run_time = verbose orelse Config.get ctxt measure_run_time
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38091
diff changeset
   387
    val command = Path.explode (getenv (fst exec) ^ "/" ^ snd exec)
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   388
    fun split_time s =
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   389
      let
42448
95b2626c75a8 tuning -- local semicolon consistency
blanchet
parents: 42447
diff changeset
   390
        val split = String.tokens (fn c => str c = "\n")
95b2626c75a8 tuning -- local semicolon consistency
blanchet
parents: 42447
diff changeset
   391
        val (output, t) = s |> split |> split_last |> apfst cat_lines
95b2626c75a8 tuning -- local semicolon consistency
blanchet
parents: 42447
diff changeset
   392
        fun as_num f = f >> (fst o read_int)
95b2626c75a8 tuning -- local semicolon consistency
blanchet
parents: 42447
diff changeset
   393
        val num = as_num (Scan.many1 Symbol.is_ascii_digit)
95b2626c75a8 tuning -- local semicolon consistency
blanchet
parents: 42447
diff changeset
   394
        val digit = Scan.one Symbol.is_ascii_digit
95b2626c75a8 tuning -- local semicolon consistency
blanchet
parents: 42447
diff changeset
   395
        val num3 = as_num (digit ::: digit ::: (digit >> single))
95b2626c75a8 tuning -- local semicolon consistency
blanchet
parents: 42447
diff changeset
   396
        val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b)
40627
becf5d5187cc renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents: 40562
diff changeset
   397
        val as_time = Scan.read Symbol.stopper time o raw_explode
42448
95b2626c75a8 tuning -- local semicolon consistency
blanchet
parents: 42447
diff changeset
   398
      in (output, as_time t) end
41313
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   399
    fun run_on prob_file =
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38091
diff changeset
   400
      case filter (curry (op =) "" o getenv o fst) (exec :: required_execs) of
38032
54448f5d151f improve detection of installed SPASS
blanchet
parents: 38028
diff changeset
   401
        (home_var, _) :: _ =>
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   402
        error ("The environment variable " ^ quote home_var ^ " is not set.")
38032
54448f5d151f improve detection of installed SPASS
blanchet
parents: 38028
diff changeset
   403
      | [] =>
54448f5d151f improve detection of installed SPASS
blanchet
parents: 38028
diff changeset
   404
        if File.exists command then
54448f5d151f improve detection of installed SPASS
blanchet
parents: 38028
diff changeset
   405
          let
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   406
            (* If slicing is disabled, we expand the last slice to fill the
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   407
               entire time available. *)
42578
1eaf4d437d4c define type system in ATP module so that ATPs can suggest a type system
blanchet
parents: 42577
diff changeset
   408
            val actual_slices = get_slices slicing (best_slices ())
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   409
            val num_actual_slices = length actual_slices
42445
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   410
            fun monomorphize_facts facts =
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   411
              let
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   412
                val repair_context =
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   413
                  Config.put SMT_Config.verbose debug
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   414
                  #> Config.put SMT_Config.monomorph_full false
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   415
                  #> Config.put SMT_Config.monomorph_limit monomorphize_limit
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   416
                val facts = facts |> map untranslated_fact
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   417
                (* pseudo-theorem involving the same constants as the subgoal *)
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   418
                val subgoal_th =
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   419
                  Logic.list_implies (hyp_ts, concl_t)
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   420
                  |> Skip_Proof.make_thm thy
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   421
                val indexed_facts =
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   422
                  (~1, subgoal_th) :: (0 upto length facts - 1 ~~ map snd facts)
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   423
              in
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   424
                SMT_Monomorph.monomorph indexed_facts (repair_context ctxt)
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   425
                |> fst |> sort (int_ord o pairself fst)
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   426
                |> filter_out (curry (op =) ~1 o fst)
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   427
                |> map (Untranslated_Fact o apfst (fst o nth facts))
c6ea64ebb8c5 fixed interaction between monomorphization and slicing for ATPs
blanchet
parents: 42444
diff changeset
   428
              end
42452
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   429
            fun run_slice blacklist
42450
2765d4fb2b9c automatically retry with full-types upon unsound proof
blanchet
parents: 42449
diff changeset
   430
                          (slice, (time_frac, (complete, default_max_relevant)))
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   431
                          time_left =
38032
54448f5d151f improve detection of installed SPASS
blanchet
parents: 38028
diff changeset
   432
              let
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   433
                val num_facts =
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   434
                  length facts |> is_none max_relevant
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   435
                                  ? Integer.min default_max_relevant
42451
a75fcd103cbb automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents: 42450
diff changeset
   436
                val facts =
a75fcd103cbb automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents: 42450
diff changeset
   437
                  facts |> take num_facts
42452
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   438
                        |> not (null blacklist)
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   439
                           ? filter_out (member (op =) blacklist o fst
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   440
                                         o untranslated_fact)
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: 42588
diff changeset
   441
                        |> polymorphism_of_type_sys type_sys <> Polymorphic
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: 42588
diff changeset
   442
                           ? monomorphize_facts
42544
75cb06eee990 reimplemented the hAPP introduction code so that it's done earlier, when the types are still available
blanchet
parents: 42542
diff changeset
   443
                        |> map (atp_translated_fact ctxt)
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   444
                val real_ms = Real.fromInt o Time.toMilliseconds
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   445
                val slice_timeout =
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   446
                  ((real_ms time_left
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   447
                    |> (if slice < num_actual_slices - 1 then
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   448
                          curry Real.min (time_frac * real_ms timeout)
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   449
                        else
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   450
                          I))
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   451
                   * 0.001) |> seconds
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   452
                val _ =
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   453
                  if verbose then
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   454
                    "ATP slice " ^ string_of_int (slice + 1) ^ " with " ^
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   455
                    string_of_int num_facts ^ " fact" ^ plural_s num_facts ^
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   456
                    " for " ^ string_from_time slice_timeout ^ "..."
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   457
                    |> Output.urgent_message
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   458
                  else
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   459
                    ()
42541
8938507b2054 move type declarations to the front, for TFF-compliance
blanchet
parents: 42523
diff changeset
   460
                val (atp_problem, pool, conjecture_offset, facts_offset,
8938507b2054 move type declarations to the front, for TFF-compliance
blanchet
parents: 42523
diff changeset
   461
                     fact_names) =
42568
7b9801a34836 no needless "fequal" proxies if "explicit_apply" is set + always have readable names
blanchet
parents: 42548
diff changeset
   462
                  prepare_atp_problem ctxt type_sys explicit_apply hyp_ts
7b9801a34836 no needless "fequal" proxies if "explicit_apply" is set + always have readable names
blanchet
parents: 42548
diff changeset
   463
                                      concl_t facts
41313
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   464
                fun weights () = atp_problem_weights atp_problem
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   465
                val core =
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   466
                  File.shell_path command ^ " " ^
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   467
                  arguments slice slice_timeout weights ^ " " ^
41313
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   468
                  File.shell_path prob_file
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   469
                val command =
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   470
                  (if measure_run_time then
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   471
                     "TIMEFORMAT='%3R'; { time " ^ core ^ " ; }"
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   472
                   else
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   473
                     "exec " ^ core) ^ " 2>&1"
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   474
                val _ =
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   475
                  atp_problem
42577
78414ec6fa4e made the format (TFF or FOF) of the TPTP problem a global argument of the problem again and have the ATPs report which formats they support
blanchet
parents: 42568
diff changeset
   476
                  |> tptp_strings_for_atp_problem hypothesis_kind format
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   477
                  |> cons ("% " ^ command ^ "\n")
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   478
                  |> File.write_list prob_file
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   479
                val conjecture_shape =
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   480
                  conjecture_offset + 1
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   481
                    upto conjecture_offset + length hyp_ts + 1
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   482
                  |> map single
38032
54448f5d151f improve detection of installed SPASS
blanchet
parents: 38028
diff changeset
   483
                val ((output, msecs), res_code) =
54448f5d151f improve detection of installed SPASS
blanchet
parents: 38028
diff changeset
   484
                  bash_output command
54448f5d151f improve detection of installed SPASS
blanchet
parents: 38028
diff changeset
   485
                  |>> (if overlord then
54448f5d151f improve detection of installed SPASS
blanchet
parents: 38028
diff changeset
   486
                         prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
54448f5d151f improve detection of installed SPASS
blanchet
parents: 38028
diff changeset
   487
                       else
54448f5d151f improve detection of installed SPASS
blanchet
parents: 38028
diff changeset
   488
                         I)
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   489
                  |>> (if measure_run_time then split_time else rpair NONE)
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42448
diff changeset
   490
                val (atp_proof, outcome) =
42060
889d767ce5f4 make Minimizer honor "verbose" and "debug" options better
blanchet
parents: 41990
diff changeset
   491
                  extract_tstplike_proof_and_outcome debug verbose complete
889d767ce5f4 make Minimizer honor "verbose" and "debug" options better
blanchet
parents: 41990
diff changeset
   492
                      res_code proof_delims known_failures output
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42448
diff changeset
   493
                  |>> atp_proof_from_tstplike_proof
42587
4fbb1de05169 fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
blanchet
parents: 42579
diff changeset
   494
                val (conjecture_shape, facts_offset, fact_names) =
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42448
diff changeset
   495
                  if is_none outcome then
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42448
diff changeset
   496
                    repair_conjecture_shape_and_fact_names output
42587
4fbb1de05169 fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
blanchet
parents: 42579
diff changeset
   497
                        conjecture_shape facts_offset fact_names
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42448
diff changeset
   498
                  else
42587
4fbb1de05169 fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
blanchet
parents: 42579
diff changeset
   499
                    (* don't bother repairing *)
4fbb1de05169 fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
blanchet
parents: 42579
diff changeset
   500
                    (conjecture_shape, facts_offset, fact_names)
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42448
diff changeset
   501
                val outcome =
42451
a75fcd103cbb automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents: 42450
diff changeset
   502
                  case outcome of
42587
4fbb1de05169 fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
blanchet
parents: 42579
diff changeset
   503
                    NONE =>
4fbb1de05169 fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
blanchet
parents: 42579
diff changeset
   504
                    if is_unsound_proof conjecture_shape facts_offset fact_names
4fbb1de05169 fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
blanchet
parents: 42579
diff changeset
   505
                                        atp_proof 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: 42588
diff changeset
   506
                      SOME (UnsoundProof (is_type_sys_virtually_sound type_sys))
42587
4fbb1de05169 fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
blanchet
parents: 42579
diff changeset
   507
                    else
4fbb1de05169 fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
blanchet
parents: 42579
diff changeset
   508
                      NONE
42451
a75fcd103cbb automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents: 42450
diff changeset
   509
                  | SOME Unprovable =>
42452
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   510
                    if null blacklist then outcome
42451
a75fcd103cbb automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents: 42450
diff changeset
   511
                    else SOME IncompleteUnprovable
a75fcd103cbb automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents: 42450
diff changeset
   512
                  | _ => outcome
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   513
              in
42541
8938507b2054 move type declarations to the front, for TFF-compliance
blanchet
parents: 42523
diff changeset
   514
                ((pool, conjecture_shape, facts_offset, fact_names),
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42448
diff changeset
   515
                 (output, msecs, atp_proof, outcome))
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   516
              end
38645
4d5bbec1a598 be more generous towards SPASS's -SOS mode
blanchet
parents: 38631
diff changeset
   517
            val timer = Timer.startRealTimer ()
42452
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   518
            fun maybe_run_slice blacklist slice
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   519
                                (result as (_, (_, msecs0, _, SOME _))) =
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   520
                let
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   521
                  val time_left = Time.- (timeout, Timer.checkRealTimer timer)
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   522
                in
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   523
                  if Time.<= (time_left, Time.zeroTime) then
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   524
                    result
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   525
                  else
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   526
                    (run_slice blacklist slice time_left
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   527
                     |> (fn (stuff, (output, msecs, atp_proof, outcome)) =>
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   528
                            (stuff, (output, int_opt_add msecs0 msecs, atp_proof,
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   529
                                     outcome))))
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   530
                end
42450
2765d4fb2b9c automatically retry with full-types upon unsound proof
blanchet
parents: 42449
diff changeset
   531
              | maybe_run_slice _ _ result = result
42452
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   532
            fun maybe_blacklist_facts_and_retry iter blacklist
42541
8938507b2054 move type declarations to the front, for TFF-compliance
blanchet
parents: 42523
diff changeset
   533
                    (result as ((_, _, facts_offset, fact_names),
42587
4fbb1de05169 fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
blanchet
parents: 42579
diff changeset
   534
                                (_, _, atp_proof, SOME (UnsoundProof false)))) =
42541
8938507b2054 move type declarations to the front, for TFF-compliance
blanchet
parents: 42523
diff changeset
   535
                (case used_facts_in_atp_proof facts_offset fact_names
8938507b2054 move type declarations to the front, for TFF-compliance
blanchet
parents: 42523
diff changeset
   536
                                              atp_proof of
42452
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   537
                   [] => result
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   538
                 | new_baddies =>
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   539
                   let val blacklist = new_baddies @ blacklist in
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   540
                     result
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   541
                     |> maybe_run_slice blacklist (List.last actual_slices)
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   542
                     |> iter < atp_blacklist_max_iters
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   543
                        ? maybe_blacklist_facts_and_retry (iter + 1) blacklist
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   544
                   end)
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   545
              | maybe_blacklist_facts_and_retry _ _ result = result
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   546
          in
42541
8938507b2054 move type declarations to the front, for TFF-compliance
blanchet
parents: 42523
diff changeset
   547
            ((Symtab.empty, [], 0, Vector.fromList []),
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42448
diff changeset
   548
             ("", SOME 0, [], SOME InternalError))
42451
a75fcd103cbb automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents: 42450
diff changeset
   549
            |> fold (maybe_run_slice []) actual_slices
a75fcd103cbb automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents: 42450
diff changeset
   550
               (* The ATP found an unsound proof? Automatically try again
a75fcd103cbb automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents: 42450
diff changeset
   551
                  without the offending facts! *)
42452
f7f796ce5d68 iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents: 42451
diff changeset
   552
            |> maybe_blacklist_facts_and_retry 0 []
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   553
          end
38032
54448f5d151f improve detection of installed SPASS
blanchet
parents: 38028
diff changeset
   554
        else
41944
b97091ae583a Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
wenzelm
parents: 41799
diff changeset
   555
          error ("Bad executable: " ^ Path.print command ^ ".")
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   556
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   557
    (* If the problem file has not been exported, remove it; otherwise, export
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   558
       the proof file too. *)
41313
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   559
    fun cleanup prob_file =
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   560
      if dest_dir = "" then try File.rm prob_file else NONE
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   561
    fun export prob_file (_, (output, _, _, _)) =
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39494
diff changeset
   562
      if dest_dir = "" then
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   563
        ()
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   564
      else
41313
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   565
        File.write (Path.explode (Path.implode prob_file ^ "_proof")) output
42541
8938507b2054 move type declarations to the front, for TFF-compliance
blanchet
parents: 42523
diff changeset
   566
    val ((pool, conjecture_shape, facts_offset, fact_names),
42449
494e4ac5b0f8 detect some unsound proofs before showing them to the user
blanchet
parents: 42448
diff changeset
   567
         (output, msecs, atp_proof, outcome)) =
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   568
      with_path cleanup export run_on problem_path_name
39492
b1172d65dd28 skip some "important" messages
blanchet
parents: 39491
diff changeset
   569
    val important_message =
42609
b5e94b70bc06 fixed random number invocation
blanchet
parents: 42593
diff changeset
   570
      if not auto andalso
b5e94b70bc06 fixed random number invocation
blanchet
parents: 42593
diff changeset
   571
         random_range 0 (atp_important_message_keep_quotient - 1) = 0 then
39492
b1172d65dd28 skip some "important" messages
blanchet
parents: 39491
diff changeset
   572
        extract_important_message output
b1172d65dd28 skip some "important" messages
blanchet
parents: 39491
diff changeset
   573
      else
b1172d65dd28 skip some "important" messages
blanchet
parents: 39491
diff changeset
   574
        ""
41742
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   575
    fun append_to_message message =
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   576
      message ^
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   577
      (if verbose then
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   578
         "\nATP real CPU time: " ^
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   579
         string_from_time (Time.fromMilliseconds (the msecs)) ^ "."
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   580
       else
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   581
         "") ^
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   582
      (if important_message <> "" then
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   583
         "\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^ important_message
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   584
       else
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   585
         "")
42593
f9d7f1331a00 use "metis", not "metisFT", to reconstruct proofs found in fully-typed mode -- "metisFT" is just too slow...
blanchet
parents: 42589
diff changeset
   586
    val isar_params =
f9d7f1331a00 use "metis", not "metisFT", to reconstruct proofs found in fully-typed mode -- "metisFT" is just too slow...
blanchet
parents: 42589
diff changeset
   587
      (ctxt, debug, type_sys, isar_shrink_factor, pool, conjecture_shape)
41742
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   588
    val metis_params =
42593
f9d7f1331a00 use "metis", not "metisFT", to reconstruct proofs found in fully-typed mode -- "metisFT" is just too slow...
blanchet
parents: 42589
diff changeset
   589
      (proof_banner auto, minimize_command, atp_proof, facts_offset, fact_names,
f9d7f1331a00 use "metis", not "metisFT", to reconstruct proofs found in fully-typed mode -- "metisFT" is just too slow...
blanchet
parents: 42589
diff changeset
   590
       goal, subgoal)
41742
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   591
    val (outcome, (message, used_facts)) =
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   592
      case outcome of
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   593
        NONE =>
41742
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   594
        (NONE, proof_text isar_proof isar_params metis_params
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   595
               |>> append_to_message)
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   596
      | SOME ProofMissing =>
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   597
        (NONE, metis_proof_text metis_params |>> append_to_message)
41744
a18e7bbca258 make minimizer verbose
blanchet
parents: 41742
diff changeset
   598
      | SOME failure => (outcome, (string_for_failure failure, []))
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   599
  in
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40202
diff changeset
   600
    {outcome = outcome, message = message, used_facts = used_facts,
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   601
     run_time_in_msecs = msecs}
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   602
  end
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   603
40669
5c316d1327d4 more precise error handling for Z3;
blanchet
parents: 40668
diff changeset
   604
(* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until
5c316d1327d4 more precise error handling for Z3;
blanchet
parents: 40668
diff changeset
   605
   these are sorted out properly in the SMT module, we have to interpret these
5c316d1327d4 more precise error handling for Z3;
blanchet
parents: 40668
diff changeset
   606
   ourselves. *)
40684
c7ba327eb58c more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents: 40669
diff changeset
   607
val remote_smt_failures =
c7ba327eb58c more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents: 40669
diff changeset
   608
  [(22, CantConnect),
c7ba327eb58c more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents: 40669
diff changeset
   609
   (2, NoLibwwwPerl)]
41222
f9783376d9b1 more precise/correct SMT error handling
blanchet
parents: 41220
diff changeset
   610
val z3_wrapper_failures =
f9783376d9b1 more precise/correct SMT error handling
blanchet
parents: 41220
diff changeset
   611
  [(10, NoRealZ3),
f9783376d9b1 more precise/correct SMT error handling
blanchet
parents: 41220
diff changeset
   612
   (11, InternalError),
f9783376d9b1 more precise/correct SMT error handling
blanchet
parents: 41220
diff changeset
   613
   (12, InternalError),
f9783376d9b1 more precise/correct SMT error handling
blanchet
parents: 41220
diff changeset
   614
   (13, InternalError)]
40684
c7ba327eb58c more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents: 40669
diff changeset
   615
val z3_failures =
41236
def0a3013554 trap one more Z3 error
blanchet
parents: 41222
diff changeset
   616
  [(101, OutOfResources),
def0a3013554 trap one more Z3 error
blanchet
parents: 41222
diff changeset
   617
   (103, MalformedInput),
41222
f9783376d9b1 more precise/correct SMT error handling
blanchet
parents: 41220
diff changeset
   618
   (110, MalformedInput)]
40684
c7ba327eb58c more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents: 40669
diff changeset
   619
val unix_failures =
c7ba327eb58c more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents: 40669
diff changeset
   620
  [(139, Crashed)]
41222
f9783376d9b1 more precise/correct SMT error handling
blanchet
parents: 41220
diff changeset
   621
val smt_failures =
41799
98b3d5ce0935 exit code 127 can mean many things -- not just (and probably not) Perl missing
blanchet
parents: 41790
diff changeset
   622
  remote_smt_failures @ z3_wrapper_failures @ z3_failures @ unix_failures
40555
de581d7da0b6 interpret SMT_Failure.Solver_Crashed correctly
blanchet
parents: 40553
diff changeset
   623
42100
062381c5f9f8 more precise failure reporting in Sledgehammer/SMT
blanchet
parents: 42061
diff changeset
   624
fun failure_from_smt_failure (SMT_Failure.Counterexample {is_real_cex, ...}) =
062381c5f9f8 more precise failure reporting in Sledgehammer/SMT
blanchet
parents: 42061
diff changeset
   625
    if is_real_cex then Unprovable else IncompleteUnprovable
41222
f9783376d9b1 more precise/correct SMT error handling
blanchet
parents: 41220
diff changeset
   626
  | failure_from_smt_failure SMT_Failure.Time_Out = TimedOut
f9783376d9b1 more precise/correct SMT error handling
blanchet
parents: 41220
diff changeset
   627
  | failure_from_smt_failure (SMT_Failure.Abnormal_Termination code) =
f9783376d9b1 more precise/correct SMT error handling
blanchet
parents: 41220
diff changeset
   628
    (case AList.lookup (op =) smt_failures code of
40684
c7ba327eb58c more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents: 40669
diff changeset
   629
       SOME failure => failure
41259
13972ced98d9 more precise error messages in "verbose" (or "debug") mode, following this morning's permission debacle
blanchet
parents: 41256
diff changeset
   630
     | NONE => UnknownError ("Abnormal termination with exit code " ^
13972ced98d9 more precise error messages in "verbose" (or "debug") mode, following this morning's permission debacle
blanchet
parents: 41256
diff changeset
   631
                             string_of_int code ^ "."))
41222
f9783376d9b1 more precise/correct SMT error handling
blanchet
parents: 41220
diff changeset
   632
  | failure_from_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources
f9783376d9b1 more precise/correct SMT error handling
blanchet
parents: 41220
diff changeset
   633
  | failure_from_smt_failure (SMT_Failure.Other_Failure msg) =
42061
71077681eaf6 let SMT errors through -- the main reason for keeping them quiet was that the SMT bridge used to suffer from internal bugs, but these have been fixed for some time now
blanchet
parents: 42060
diff changeset
   634
    UnknownError msg
40063
d086e3699e78 bring ATPs and SMT solvers more in line with each other
blanchet
parents: 40062
diff changeset
   635
40698
8a3f7ea91370 cosmetics
blanchet
parents: 40693
diff changeset
   636
(* FUDGE *)
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   637
val smt_max_slices = Unsynchronized.ref 8
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   638
val smt_slice_fact_frac = Unsynchronized.ref 0.5
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   639
val smt_slice_time_frac = Unsynchronized.ref 0.5
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   640
val smt_slice_min_secs = Unsynchronized.ref 5
40409
3642dc3b72e8 invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents: 40370
diff changeset
   641
42444
8e5438dc70bb cleanup: get rid of "may_slice" arguments without changing semantics
blanchet
parents: 42443
diff changeset
   642
fun smt_filter_loop name ({debug, verbose, overlord, monomorphize_limit,
8e5438dc70bb cleanup: get rid of "may_slice" arguments without changing semantics
blanchet
parents: 42443
diff changeset
   643
                           timeout, slicing, ...} : params)
41741
839d1488045f renamed field
blanchet
parents: 41727
diff changeset
   644
                    state i smt_filter =
40409
3642dc3b72e8 invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents: 40370
diff changeset
   645
  let
40553
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   646
    val ctxt = Proof.context_of state
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   647
    val max_slices = if slicing then !smt_max_slices else 1
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   648
    val repair_context =
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41336
diff changeset
   649
      select_smt_solver name
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   650
      #> Config.put SMT_Config.verbose debug
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   651
      #> (if overlord then
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   652
            Config.put SMT_Config.debug_files
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   653
                       (overlord_file_location_for_prover name
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41336
diff changeset
   654
                        |> (fn (path, name) => path ^ "/" ^ name))
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   655
          else
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   656
            I)
41335
66edbd0f7a2e added "smt_triggers" option to experiment with triggers in Sledgehammer;
blanchet
parents: 41318
diff changeset
   657
      #> Config.put SMT_Config.infer_triggers (!smt_triggers)
42193
8eed749527b6 remove workaround 8f25605e646c, which is no longer necessary thanks to 173b0f488428
blanchet
parents: 42190
diff changeset
   658
      #> Config.put SMT_Config.monomorph_full false
42180
a6c141925a8a added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents: 42100
diff changeset
   659
      #> Config.put SMT_Config.monomorph_limit monomorphize_limit
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   660
    val state = state |> Proof.map_context repair_context
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   661
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   662
    fun do_slice timeout slice outcome0 time_so_far facts =
40553
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   663
      let
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   664
        val timer = Timer.startRealTimer ()
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   665
        val ms = timeout |> Time.toMilliseconds
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   666
        val slice_timeout =
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   667
          if slice < max_slices then
41169
95167879f675 clean up fudge factors a little bit
blanchet
parents: 41168
diff changeset
   668
            Int.min (ms,
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   669
                Int.max (1000 * !smt_slice_min_secs,
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   670
                    Real.ceil (!smt_slice_time_frac * Real.fromInt ms)))
40553
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   671
            |> Time.fromMilliseconds
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   672
          else
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   673
            timeout
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   674
        val num_facts = length facts
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   675
        val _ =
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   676
          if verbose then
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   677
            "SMT slice with " ^ string_of_int num_facts ^ " fact" ^
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   678
            plural_s num_facts ^ " for " ^ string_from_time slice_timeout ^
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   679
            "..."
40553
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   680
            |> Output.urgent_message
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   681
          else
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   682
            ()
41168
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
   683
        val birth = Timer.checkRealTimer timer
41171
043f8dc3b51f facilitate debugging
blanchet
parents: 41169
diff changeset
   684
        val _ =
41211
1e2e16bc0077 no need to do a super-duper atomization if Metis fails afterwards anyway
blanchet
parents: 41209
diff changeset
   685
          if debug then Output.urgent_message "Invoking SMT solver..." else ()
41209
91fab0d3553b robustly handle SMT exceptions in Sledgehammer
blanchet
parents: 41208
diff changeset
   686
        val (outcome, used_facts) =
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   687
          (case (slice, smt_filter) of
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41336
diff changeset
   688
             (1, SOME head) => head |> apsnd (apsnd repair_context)
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41336
diff changeset
   689
           | _ => SMT_Solver.smt_filter_preprocess state facts i)
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   690
          |> SMT_Solver.smt_filter_apply slice_timeout
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41236
diff changeset
   691
          |> (fn {outcome, used_facts} => (outcome, used_facts))
41209
91fab0d3553b robustly handle SMT exceptions in Sledgehammer
blanchet
parents: 41208
diff changeset
   692
          handle exn => if Exn.is_interrupt exn then
91fab0d3553b robustly handle SMT exceptions in Sledgehammer
blanchet
parents: 41208
diff changeset
   693
                          reraise exn
91fab0d3553b robustly handle SMT exceptions in Sledgehammer
blanchet
parents: 41208
diff changeset
   694
                        else
42061
71077681eaf6 let SMT errors through -- the main reason for keeping them quiet was that the SMT bridge used to suffer from internal bugs, but these have been fixed for some time now
blanchet
parents: 42060
diff changeset
   695
                          (ML_Compiler.exn_message exn
41209
91fab0d3553b robustly handle SMT exceptions in Sledgehammer
blanchet
parents: 41208
diff changeset
   696
                           |> SMT_Failure.Other_Failure |> SOME, [])
41168
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
   697
        val death = Timer.checkRealTimer timer
40553
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   698
        val _ =
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   699
          if verbose andalso is_some outcome then
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   700
            "SMT outcome: " ^ SMT_Failure.string_of_failure ctxt (the outcome)
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   701
            |> Output.urgent_message
41171
043f8dc3b51f facilitate debugging
blanchet
parents: 41169
diff changeset
   702
          else if debug then
043f8dc3b51f facilitate debugging
blanchet
parents: 41169
diff changeset
   703
            Output.urgent_message "SMT solver returned."
40553
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   704
          else
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   705
            ()
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   706
        val outcome0 = if is_none outcome0 then SOME outcome else outcome0
41168
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
   707
        val time_so_far = Time.+ (time_so_far, Time.- (death, birth))
40553
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   708
        val too_many_facts_perhaps =
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   709
          case outcome of
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   710
            NONE => false
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   711
          | SOME (SMT_Failure.Counterexample _) => false
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   712
          | SOME SMT_Failure.Time_Out => slice_timeout <> timeout
40561
0125cbb5d3c7 renamed SMT failure: Abnormal_Termination is indeed more appropriate than Solver_Crashed
boehmes
parents: 40558
diff changeset
   713
          | SOME (SMT_Failure.Abnormal_Termination code) =>
40553
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   714
            (if verbose then
40555
de581d7da0b6 interpret SMT_Failure.Solver_Crashed correctly
blanchet
parents: 40553
diff changeset
   715
               "The SMT solver invoked with " ^ string_of_int num_facts ^
40692
7fa054c3f810 make "debug" mode of Sledgehammer/SMT more liberal
blanchet
parents: 40684
diff changeset
   716
               " fact" ^ plural_s num_facts ^ " terminated abnormally with \
7fa054c3f810 make "debug" mode of Sledgehammer/SMT more liberal
blanchet
parents: 40684
diff changeset
   717
               \exit code " ^ string_of_int code ^ "."
7fa054c3f810 make "debug" mode of Sledgehammer/SMT more liberal
blanchet
parents: 40684
diff changeset
   718
               |> warning
40553
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   719
             else
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   720
               ();
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   721
             true (* kind of *))
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   722
          | SOME SMT_Failure.Out_Of_Memory => true
41211
1e2e16bc0077 no need to do a super-duper atomization if Metis fails afterwards anyway
blanchet
parents: 41209
diff changeset
   723
          | SOME (SMT_Failure.Other_Failure _) => true
40553
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   724
        val timeout = Time.- (timeout, Timer.checkRealTimer timer)
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   725
      in
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   726
        if too_many_facts_perhaps andalso slice < max_slices andalso
40553
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   727
           num_facts > 0 andalso Time.> (timeout, Time.zeroTime) then
41169
95167879f675 clean up fudge factors a little bit
blanchet
parents: 41168
diff changeset
   728
          let
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   729
            val n = Real.ceil (!smt_slice_fact_frac * Real.fromInt num_facts)
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   730
          in
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   731
            do_slice timeout (slice + 1) outcome0 time_so_far (take n facts)
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   732
          end
40553
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   733
        else
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40471
diff changeset
   734
          {outcome = if is_none outcome then NONE else the outcome0,
41168
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
   735
           used_facts = used_facts,
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
   736
           run_time_in_msecs = SOME (Time.toMilliseconds time_so_far)}
40409
3642dc3b72e8 invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents: 40370
diff changeset
   737
      end
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   738
  in do_slice timeout 1 NONE Time.zeroTime end
40409
3642dc3b72e8 invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents: 40370
diff changeset
   739
40666
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   740
(* taken from "Mirabelle" and generalized *)
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   741
fun can_apply timeout tac state i =
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   742
  let
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   743
    val {context = ctxt, facts, goal} = Proof.goal state
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   744
    val full_tac = Method.insert_tac facts i THEN tac ctxt i
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   745
  in
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   746
    case try (TimeLimit.timeLimit timeout (Seq.pull o full_tac)) goal of
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   747
      SOME (SOME _) => true
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   748
    | _ => false
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   749
  end
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   750
41152
4a7410be4dfc crank up Metis's timeout for SMT solvers, since users love Metis
blanchet
parents: 41151
diff changeset
   751
val smt_metis_timeout = seconds 1.0
40666
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   752
40693
a4171afcc3c4 set Metis option on correct context, lest it be ignored
blanchet
parents: 40692
diff changeset
   753
fun can_apply_metis debug state i ths =
a4171afcc3c4 set Metis option on correct context, lest it be ignored
blanchet
parents: 40692
diff changeset
   754
  can_apply smt_metis_timeout
a4171afcc3c4 set Metis option on correct context, lest it be ignored
blanchet
parents: 40692
diff changeset
   755
            (Config.put Metis_Tactics.verbose debug
a4171afcc3c4 set Metis option on correct context, lest it be ignored
blanchet
parents: 40692
diff changeset
   756
             #> (fn ctxt => Metis_Tactics.metis_tac ctxt ths)) state i
40666
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   757
42444
8e5438dc70bb cleanup: get rid of "may_slice" arguments without changing semantics
blanchet
parents: 42443
diff changeset
   758
fun run_smt_solver auto name (params as {debug, verbose, ...}) minimize_command
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   759
                   ({state, subgoal, subgoal_count, facts, smt_filter, ...}
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   760
                    : prover_problem) =
36379
20ef039bccff make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents: 36373
diff changeset
   761
  let
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41241
diff changeset
   762
    val ctxt = Proof.context_of state
41256
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   763
    val thy = Proof.theory_of state
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   764
    val num_facts = length facts
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   765
    val facts = facts ~~ (0 upto num_facts - 1)
0e7d45cc005f put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents: 41242
diff changeset
   766
                |> map (smt_weighted_fact thy num_facts)
40181
3788b7adab36 integrated "smt" proof method with Sledgehammer
blanchet
parents: 40132
diff changeset
   767
    val {outcome, used_facts, run_time_in_msecs} =
42444
8e5438dc70bb cleanup: get rid of "may_slice" arguments without changing semantics
blanchet
parents: 42443
diff changeset
   768
      smt_filter_loop name params state subgoal smt_filter facts
40723
a82badd0e6ef put facts found by SMT solver in alphabetical order and omit chained facts, as was done already for ATP proofs
blanchet
parents: 40701
diff changeset
   769
    val (chained_lemmas, other_lemmas) = split_used_facts (map fst used_facts)
41222
f9783376d9b1 more precise/correct SMT error handling
blanchet
parents: 41220
diff changeset
   770
    val outcome = outcome |> Option.map failure_from_smt_failure
40063
d086e3699e78 bring ATPs and SMT solvers more in line with each other
blanchet
parents: 40062
diff changeset
   771
    val message =
40184
91b4b73dbafb proper error handling for SMT solvers in Sledgehammer
blanchet
parents: 40181
diff changeset
   772
      case outcome of
91b4b73dbafb proper error handling for SMT solvers in Sledgehammer
blanchet
parents: 40181
diff changeset
   773
        NONE =>
40666
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   774
        let
41151
d58157bb1ae7 generate a "using [[smt_solver = ...]]" command if a proof is found by another SMT solver than the current one, to ensure it's also used for reconstruction
blanchet
parents: 41140
diff changeset
   775
          val (method, settings) =
40693
a4171afcc3c4 set Metis option on correct context, lest it be ignored
blanchet
parents: 40692
diff changeset
   776
            if can_apply_metis debug state subgoal (map snd used_facts) then
41151
d58157bb1ae7 generate a "using [[smt_solver = ...]]" command if a proof is found by another SMT solver than the current one, to ensure it's also used for reconstruction
blanchet
parents: 41140
diff changeset
   777
              ("metis", "")
40693
a4171afcc3c4 set Metis option on correct context, lest it be ignored
blanchet
parents: 40692
diff changeset
   778
            else
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41336
diff changeset
   779
              ("smt", if name = SMT_Solver.solver_name_of ctxt then ""
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41336
diff changeset
   780
                      else "smt_solver = " ^ maybe_quote name)
40666
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   781
        in
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   782
          try_command_line (proof_banner auto)
41151
d58157bb1ae7 generate a "using [[smt_solver = ...]]" command if a proof is found by another SMT solver than the current one, to ensure it's also used for reconstruction
blanchet
parents: 41140
diff changeset
   783
              (apply_on_subgoal settings subgoal subgoal_count ^
d58157bb1ae7 generate a "using [[smt_solver = ...]]" command if a proof is found by another SMT solver than the current one, to ensure it's also used for reconstruction
blanchet
parents: 41140
diff changeset
   784
               command_call method (map fst other_lemmas)) ^
40723
a82badd0e6ef put facts found by SMT solver in alphabetical order and omit chained facts, as was done already for ATP proofs
blanchet
parents: 40701
diff changeset
   785
          minimize_line minimize_command
41168
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
   786
                        (map fst (other_lemmas @ chained_lemmas)) ^
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
   787
          (if verbose then
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
   788
             "\nSMT solver real CPU time: " ^
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
   789
             string_from_time (Time.fromMilliseconds (the run_time_in_msecs)) ^
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
   790
             "."
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
   791
           else
f6f1ffd51d87 added weights to SMT problems
blanchet
parents: 41159
diff changeset
   792
             "")
40666
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   793
        end
41744
a18e7bbca258 make minimizer verbose
blanchet
parents: 41742
diff changeset
   794
      | SOME failure => string_for_failure failure
40063
d086e3699e78 bring ATPs and SMT solvers more in line with each other
blanchet
parents: 40062
diff changeset
   795
  in
40666
8db6c2b1591d try Metis to reconstruct SMT proofs, to increase success rate and reduce dependency on (often remote) SMT solvers or certificates
blanchet
parents: 40627
diff changeset
   796
    {outcome = outcome, used_facts = map fst used_facts,
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40202
diff changeset
   797
     run_time_in_msecs = run_time_in_msecs, message = message}
40063
d086e3699e78 bring ATPs and SMT solvers more in line with each other
blanchet
parents: 40062
diff changeset
   798
  end
d086e3699e78 bring ATPs and SMT solvers more in line with each other
blanchet
parents: 40062
diff changeset
   799
42444
8e5438dc70bb cleanup: get rid of "may_slice" arguments without changing semantics
blanchet
parents: 42443
diff changeset
   800
fun get_prover ctxt auto name =
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42193
diff changeset
   801
  let val thy = Proof_Context.theory_of ctxt in
40941
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40723
diff changeset
   802
    if is_smt_prover ctxt name then
42444
8e5438dc70bb cleanup: get rid of "may_slice" arguments without changing semantics
blanchet
parents: 42443
diff changeset
   803
      run_smt_solver auto name
41727
ab3f6d76fb23 available_provers ~> supported_provers (for clarity)
blanchet
parents: 41723
diff changeset
   804
    else if member (op =) (supported_atps thy) name then
42444
8e5438dc70bb cleanup: get rid of "may_slice" arguments without changing semantics
blanchet
parents: 42443
diff changeset
   805
      run_atp auto name (get_atp thy name)
40941
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40723
diff changeset
   806
    else
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40723
diff changeset
   807
      error ("No such prover: " ^ name ^ ".")
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40723
diff changeset
   808
  end
40063
d086e3699e78 bring ATPs and SMT solvers more in line with each other
blanchet
parents: 40062
diff changeset
   809
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   810
val setup =
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   811
  dest_dir_setup
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   812
  #> problem_prefix_setup
39003
c2aebd79981f run relevance filter in a thread, to avoid blocking
blanchet
parents: 39000
diff changeset
   813
  #> measure_run_time_setup
38023
962b0a7f544b more refactoring
blanchet
parents: 38021
diff changeset
   814
28582
c269a3045fdf info: back to plain printing;
wenzelm
parents: 28571
diff changeset
   815
end;