src/HOL/Tools/Sledgehammer/sledgehammer_prover_smt.ML
author wenzelm
Sat, 02 Apr 2016 23:29:05 +0200
changeset 62826 eb94e570c1a4
parent 62735 23de054397e5
child 63692 1bc4bc2c9fd1
permissions -rw-r--r--
prefer infix operations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_prover_smt.ML
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
     2
    Author:     Fabian Immler, TU Muenchen
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
     3
    Author:     Makarius
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
     4
    Author:     Jasmin Blanchette, TU Muenchen
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
     5
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
     6
SMT solvers as Sledgehammer provers.
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
     7
*)
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
     8
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
     9
signature SLEDGEHAMMER_PROVER_SMT =
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    10
sig
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    11
  type stature = ATP_Problem_Generate.stature
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    12
  type mode = Sledgehammer_Prover.mode
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    13
  type prover = Sledgehammer_Prover.prover
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    14
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    15
  val smt_builtins : bool Config.T
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    16
  val smt_triggers : bool Config.T
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    17
  val smt_max_slices : int Config.T
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    18
  val smt_slice_fact_frac : real Config.T
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    19
  val smt_slice_time_frac : real Config.T
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    20
  val smt_slice_min_secs : int Config.T
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    21
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    22
  val is_smt_prover : Proof.context -> string -> bool
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    23
  val run_smt_solver : mode -> string -> prover
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    24
end;
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    25
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    26
structure Sledgehammer_Prover_SMT : SLEDGEHAMMER_PROVER_SMT =
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    27
struct
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    28
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    29
open ATP_Util
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    30
open ATP_Proof
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    31
open ATP_Systems
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    32
open ATP_Problem_Generate
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    33
open ATP_Proof_Reconstruct
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    34
open Sledgehammer_Util
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    35
open Sledgehammer_Proof_Methods
56083
b5d1d9c60341 have Sledgehammer generate Isar proofs from Z3 proofs
blanchet
parents: 56082
diff changeset
    36
open Sledgehammer_Isar
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    37
open Sledgehammer_Prover
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    38
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    39
val smt_builtins = Attrib.setup_config_bool @{binding sledgehammer_smt_builtins} (K true)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    40
val smt_triggers = Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    41
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    42
val is_smt_prover = member (op =) o SMT_Config.available_solvers_of
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    43
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    44
(* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until these are sorted out
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    45
   properly in the SMT module, we must interpret these here. *)
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    46
val z3_failures =
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    47
  [(101, OutOfResources),
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    48
   (103, MalformedInput),
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    49
   (110, MalformedInput),
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    50
   (112, TimedOut)]
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    51
val unix_failures =
59019
0c58b5cf989a other way of crashing (with CVC4)
blanchet
parents: 58843
diff changeset
    52
  [(134, Crashed),
0c58b5cf989a other way of crashing (with CVC4)
blanchet
parents: 58843
diff changeset
    53
   (138, Crashed),
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    54
   (139, Crashed)]
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    55
val smt_failures = z3_failures @ unix_failures
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    56
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    57
fun failure_of_smt_failure (SMT_Failure.Counterexample genuine) =
57158
f028d93798e6 simplified counterexample handling
blanchet
parents: 57078
diff changeset
    58
    if genuine then Unprovable else GaveUp
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    59
  | failure_of_smt_failure SMT_Failure.Time_Out = TimedOut
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    60
  | failure_of_smt_failure (SMT_Failure.Abnormal_Termination code) =
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    61
    (case AList.lookup (op =) smt_failures code of
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    62
      SOME failure => failure
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    63
    | NONE => UnknownError ("Abnormal termination with exit code " ^ string_of_int code ^ "."))
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    64
  | failure_of_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    65
  | failure_of_smt_failure (SMT_Failure.Other_Failure s) = UnknownError s
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    66
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    67
(* FUDGE *)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    68
val smt_max_slices = Attrib.setup_config_int @{binding sledgehammer_smt_max_slices} (K 8)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    69
val smt_slice_fact_frac =
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    70
  Attrib.setup_config_real @{binding sledgehammer_smt_slice_fact_frac} (K 0.667)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    71
val smt_slice_time_frac =
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    72
  Attrib.setup_config_real @{binding sledgehammer_smt_slice_time_frac} (K 0.333)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    73
val smt_slice_min_secs = Attrib.setup_config_int @{binding sledgehammer_smt_slice_min_secs} (K 3)
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    74
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    75
val is_boring_builtin_typ =
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    76
  not o exists_subtype (member (op =) [@{typ nat}, @{typ int}, HOLogic.realT])
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    77
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    78
fun smt_filter_loop name ({debug, overlord, max_mono_iters, max_new_mono_instances, timeout, slice,
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    79
      ...} : params) state goal i =
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    80
  let
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    81
    fun repair_context ctxt =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    82
      ctxt |> Context.proof_map (SMT_Config.select_solver name)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    83
           |> Config.put SMT_Config.verbose debug
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    84
           |> (if overlord then
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    85
                 Config.put SMT_Config.debug_files
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    86
                   (overlord_file_location_of_prover name |> (fn (path, name) => path ^ "/" ^ name))
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    87
               else
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    88
                 I)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    89
           |> Config.put SMT_Config.infer_triggers (Config.get ctxt smt_triggers)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    90
           |> not (Config.get ctxt smt_builtins)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    91
              ? (SMT_Builtin.filter_builtins is_boring_builtin_typ
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    92
                 #> Config.put SMT_Systems.z3_extensions false)
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    93
           |> repair_monomorph_context max_mono_iters default_max_mono_iters max_new_mono_instances
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    94
                default_max_new_mono_instances
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    95
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    96
    val state = Proof.map_context (repair_context) state
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    97
    val ctxt = Proof.context_of state
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
    98
    val max_slices = if slice then Config.get ctxt smt_max_slices else 1
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
    99
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 57159
diff changeset
   100
    fun do_slice timeout slice outcome0 time_so_far (factss as (fact_filter, facts) :: _) =
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   101
      let
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   102
        val timer = Timer.startRealTimer ()
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   103
        val slice_timeout =
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   104
          if slice < max_slices then
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   105
            let val ms = Time.toMilliseconds timeout in
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   106
              Int.min (ms, Int.max (1000 * Config.get ctxt smt_slice_min_secs,
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   107
                Real.ceil (Config.get ctxt smt_slice_time_frac * Real.fromInt ms)))
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   108
              |> Time.fromMilliseconds
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   109
            end
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   110
          else
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   111
            timeout
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 57159
diff changeset
   112
        val num_facts = length facts
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   113
        val _ =
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   114
          if debug then
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   115
            quote name ^ " slice " ^ string_of_int slice ^ " with " ^ string_of_int num_facts ^
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   116
            " fact" ^ plural_s num_facts ^ " for " ^ string_of_time slice_timeout
58843
521cea5fa777 discontinued obsolete Output.urgent_message;
wenzelm
parents: 58498
diff changeset
   117
            |> writeln
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   118
          else
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   119
            ()
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   120
        val birth = Timer.checkRealTimer timer
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   121
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
   122
        val filter_result as {outcome, ...} =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   123
          SMT_Solver.smt_filter ctxt goal facts i slice_timeout
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   124
          handle exn =>
56094
2adbc6e4cd8f let exception pass through in debug mode
blanchet
parents: 56090
diff changeset
   125
            if Exn.is_interrupt exn orelse debug then
62505
9e2a65912111 clarified modules;
wenzelm
parents: 60924
diff changeset
   126
              Exn.reraise exn
56083
b5d1d9c60341 have Sledgehammer generate Isar proofs from Z3 proofs
blanchet
parents: 56082
diff changeset
   127
            else
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   128
              {outcome = SOME (SMT_Failure.Other_Failure (Runtime.exn_message exn)),
60201
90e88e521e0e made CVC4 support work also without unsat cores
blanchet
parents: 59019
diff changeset
   129
               fact_ids = NONE, atp_proof = K []}
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   130
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   131
        val death = Timer.checkRealTimer timer
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   132
        val outcome0 = if is_none outcome0 then SOME outcome else outcome0
62826
eb94e570c1a4 prefer infix operations;
wenzelm
parents: 62735
diff changeset
   133
        val time_so_far = time_so_far + (death - birth)
eb94e570c1a4 prefer infix operations;
wenzelm
parents: 62735
diff changeset
   134
        val timeout = timeout - Timer.checkRealTimer timer
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   135
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   136
        val too_many_facts_perhaps =
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   137
          (case outcome of
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   138
            NONE => false
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   139
          | SOME (SMT_Failure.Counterexample _) => false
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   140
          | SOME SMT_Failure.Time_Out => slice_timeout <> timeout
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   141
          | SOME (SMT_Failure.Abnormal_Termination _) => true (* kind of *)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   142
          | SOME SMT_Failure.Out_Of_Memory => true
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   143
          | SOME (SMT_Failure.Other_Failure _) => true)
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   144
      in
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   145
        if too_many_facts_perhaps andalso slice < max_slices andalso num_facts > 0 andalso
62826
eb94e570c1a4 prefer infix operations;
wenzelm
parents: 62735
diff changeset
   146
           timeout > Time.zeroTime then
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   147
          let
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   148
            val new_num_facts =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   149
              Real.ceil (Config.get ctxt smt_slice_fact_frac * Real.fromInt num_facts)
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 57159
diff changeset
   150
            val factss as (new_fact_filter, _) :: _ =
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 57159
diff changeset
   151
              factss
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   152
              |> (fn (x :: xs) => xs @ [x])
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   153
              |> app_hd (apsnd (take new_num_facts))
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   154
            val show_filter = fact_filter <> new_fact_filter
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   155
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   156
            fun num_of_facts fact_filter num_facts =
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   157
              string_of_int num_facts ^ (if show_filter then " " ^ quote fact_filter else "") ^
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   158
              " fact" ^ plural_s num_facts
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   159
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   160
            val _ =
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   161
              if debug then
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   162
                quote name ^ " invoked with " ^
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   163
                num_of_facts fact_filter num_facts ^ ": " ^
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   164
                string_of_atp_failure (failure_of_smt_failure (the outcome)) ^
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   165
                " Retrying with " ^ num_of_facts new_fact_filter new_num_facts ^
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   166
                "..."
58843
521cea5fa777 discontinued obsolete Output.urgent_message;
wenzelm
parents: 58498
diff changeset
   167
                |> writeln
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   168
              else
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   169
                ()
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   170
          in
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 57159
diff changeset
   171
            do_slice timeout (slice + 1) outcome0 time_so_far factss
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   172
          end
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   173
        else
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
   174
          {outcome = if is_none outcome then NONE else the outcome0, filter_result = filter_result,
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 57159
diff changeset
   175
           used_from = facts, run_time = time_so_far}
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   176
      end
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   177
  in
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   178
    do_slice timeout 1 NONE Time.zeroTime
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   179
  end
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   180
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   181
fun run_smt_solver mode name (params as {debug, verbose, isar_proofs, compress, try0, smt_proofs,
57245
f6bf6d5341ee renamed Sledgehammer options
blanchet
parents: 57243
diff changeset
   182
      minimize, preplay_timeout, ...})
62735
23de054397e5 early warning when Sledgehammer finds a proof
blanchet
parents: 62505
diff changeset
   183
    ({state, goal, subgoal, subgoal_count, factss, found_proof, ...} : prover_problem) =
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   184
  let
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   185
    val thy = Proof.theory_of state
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   186
    val ctxt = Proof.context_of state
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   187
57243
8c261f0a9b32 reintroduced vital 'Thm.transfer'
blanchet
parents: 57165
diff changeset
   188
    val factss = map (apsnd (map (apsnd (Thm.transfer thy)))) factss
8c261f0a9b32 reintroduced vital 'Thm.transfer'
blanchet
parents: 57165
diff changeset
   189
57159
24cbdebba35a refactored Z3 to Isar proof construction code
blanchet
parents: 57158
diff changeset
   190
    val {outcome, filter_result = {fact_ids, atp_proof, ...}, used_from, run_time} =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   191
      smt_filter_loop name params state goal subgoal factss
60201
90e88e521e0e made CVC4 support work also without unsat cores
blanchet
parents: 59019
diff changeset
   192
    val used_facts =
90e88e521e0e made CVC4 support work also without unsat cores
blanchet
parents: 59019
diff changeset
   193
      (case fact_ids of
90e88e521e0e made CVC4 support work also without unsat cores
blanchet
parents: 59019
diff changeset
   194
        NONE => map fst used_from
60924
610794dff23c tuned signature, in accordance to sortBy in Scala;
wenzelm
parents: 60201
diff changeset
   195
      | SOME ids => sort_by fst (map (fst o snd) ids))
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57776
diff changeset
   196
    val outcome = Option.map failure_of_smt_failure outcome
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   197
57738
25d1495e6641 eliminated needlessly complex message tail
blanchet
parents: 57735
diff changeset
   198
    val (preferred_methss, message) =
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   199
      (case outcome of
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   200
        NONE =>
57734
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57732
diff changeset
   201
        let
62735
23de054397e5 early warning when Sledgehammer finds a proof
blanchet
parents: 62505
diff changeset
   202
          val _ = found_proof ();
58498
59bdd4f57255 tuned output in case of one-liner failure
blanchet
parents: 58061
diff changeset
   203
          val smt_method = smt_proofs <> SOME false
57734
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57732
diff changeset
   204
          val preferred_methss =
58498
59bdd4f57255 tuned output in case of one-liner failure
blanchet
parents: 58061
diff changeset
   205
            (if smt_method then SMT_Method else Metis_Method (NONE, NONE),
59bdd4f57255 tuned output in case of one-liner failure
blanchet
parents: 58061
diff changeset
   206
             bunches_of_proof_methods try0 smt_method false liftingN)
57734
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57732
diff changeset
   207
        in
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57732
diff changeset
   208
          (preferred_methss,
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57732
diff changeset
   209
           fn preplay =>
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57732
diff changeset
   210
             let
58843
521cea5fa777 discontinued obsolete Output.urgent_message;
wenzelm
parents: 58498
diff changeset
   211
               val _ = if verbose then writeln "Generating proof text..." else ()
57723
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   212
57734
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57732
diff changeset
   213
               fun isar_params () =
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57732
diff changeset
   214
                 (verbose, (NONE, NONE), preplay_timeout, compress, try0, minimize, atp_proof (),
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57732
diff changeset
   215
                  goal)
57056
8b2283566f6e properly reconstruct helpers in Z3 proofs
blanchet
parents: 57054
diff changeset
   216
57750
670cbec816b9 restored a bit of laziness
blanchet
parents: 57742
diff changeset
   217
               val one_line_params = (preplay (), proof_banner mode name, subgoal, subgoal_count)
57734
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57732
diff changeset
   218
               val num_chained = length (#facts (Proof.goal state))
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57732
diff changeset
   219
             in
57738
25d1495e6641 eliminated needlessly complex message tail
blanchet
parents: 57735
diff changeset
   220
               proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained
25d1495e6641 eliminated needlessly complex message tail
blanchet
parents: 57735
diff changeset
   221
                 one_line_params
25d1495e6641 eliminated needlessly complex message tail
blanchet
parents: 57735
diff changeset
   222
             end)
57734
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57732
diff changeset
   223
        end
57738
25d1495e6641 eliminated needlessly complex message tail
blanchet
parents: 57735
diff changeset
   224
      | SOME failure => ((Auto_Method (* dummy *), []), fn _ => string_of_atp_failure failure))
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   225
  in
57734
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57732
diff changeset
   226
    {outcome = outcome, used_facts = used_facts, used_from = used_from,
57738
25d1495e6641 eliminated needlessly complex message tail
blanchet
parents: 57735
diff changeset
   227
     preferred_methss = preferred_methss, run_time = run_time, message = message}
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   228
  end
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   229
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents:
diff changeset
   230
end;