src/HOL/Tools/Sledgehammer/sledgehammer_prover.ML
changeset 55201 1ee776da8da7
parent 55183 17ec4a29ef71
child 55202 824c48a539c9
equal deleted inserted replaced
55200:777328c9f1ea 55201:1ee776da8da7
       
     1 (*  Title:      HOL/Tools/Sledgehammer/sledgehammer_prover.ML
       
     2     Author:     Fabian Immler, TU Muenchen
       
     3     Author:     Makarius
       
     4     Author:     Jasmin Blanchette, TU Muenchen
       
     5 
       
     6 Generic prover abstraction for Sledgehammer.
       
     7 *)
       
     8 
       
     9 signature SLEDGEHAMMER_PROVER =
       
    10 sig
       
    11   type atp_failure = ATP_Proof.atp_failure
       
    12   type stature = ATP_Problem_Generate.stature
       
    13   type type_enc = ATP_Problem_Generate.type_enc
       
    14   type fact = Sledgehammer_Fact.fact
       
    15   type reconstructor = Sledgehammer_Reconstructor.reconstructor
       
    16   type play_outcome = Sledgehammer_Reconstructor.play_outcome
       
    17   type minimize_command = Sledgehammer_Reconstructor.minimize_command
       
    18 
       
    19   datatype mode = Auto_Try | Try | Normal | MaSh | Auto_Minimize | Minimize
       
    20 
       
    21   type params =
       
    22     {debug : bool,
       
    23      verbose : bool,
       
    24      overlord : bool,
       
    25      spy : bool,
       
    26      blocking : bool,
       
    27      provers : string list,
       
    28      type_enc : string option,
       
    29      strict : bool,
       
    30      lam_trans : string option,
       
    31      uncurried_aliases : bool option,
       
    32      learn : bool,
       
    33      fact_filter : string option,
       
    34      max_facts : int option,
       
    35      fact_thresholds : real * real,
       
    36      max_mono_iters : int option,
       
    37      max_new_mono_instances : int option,
       
    38      isar_proofs : bool option,
       
    39      compress_isar : real,
       
    40      try0_isar : bool,
       
    41      slice : bool,
       
    42      minimize : bool option,
       
    43      timeout : Time.time,
       
    44      preplay_timeout : Time.time,
       
    45      expect : string}
       
    46 
       
    47   type prover_problem =
       
    48     {comment : string,
       
    49      state : Proof.state,
       
    50      goal : thm,
       
    51      subgoal : int,
       
    52      subgoal_count : int,
       
    53      factss : (string * fact list) list}
       
    54 
       
    55   type prover_result =
       
    56     {outcome : atp_failure option,
       
    57      used_facts : (string * stature) list,
       
    58      used_from : fact list,
       
    59      run_time : Time.time,
       
    60      preplay : (reconstructor * play_outcome) Lazy.lazy,
       
    61      message : reconstructor * play_outcome -> string,
       
    62      message_tail : string}
       
    63 
       
    64   type prover =
       
    65     params -> ((string * string list) list -> string -> minimize_command)
       
    66     -> prover_problem -> prover_result
       
    67 
       
    68   val dest_dir : string Config.T
       
    69   val problem_prefix : string Config.T
       
    70   val completish : bool Config.T
       
    71   val atp_full_names : bool Config.T
       
    72   val smt_builtins : bool Config.T
       
    73   val smt_triggers : bool Config.T
       
    74   val smt_weights : bool Config.T
       
    75   val smt_weight_min_facts : int Config.T
       
    76   val smt_min_weight : int Config.T
       
    77   val smt_max_weight : int Config.T
       
    78   val smt_max_weight_index : int Config.T
       
    79   val smt_weight_curve : (int -> int) Unsynchronized.ref
       
    80   val smt_max_slices : int Config.T
       
    81   val smt_slice_fact_frac : real Config.T
       
    82   val smt_slice_time_frac : real Config.T
       
    83   val smt_slice_min_secs : int Config.T
       
    84   val SledgehammerN : string
       
    85   val plain_metis : reconstructor
       
    86   val select_smt_solver : string -> Proof.context -> Proof.context
       
    87   val extract_reconstructor : params -> reconstructor -> string * (string * string list) list
       
    88   val is_reconstructor : string -> bool
       
    89   val is_atp : theory -> string -> bool
       
    90   val is_smt_prover : Proof.context -> string -> bool
       
    91   val is_ho_atp: Proof.context -> string -> bool
       
    92   val is_unit_equational_atp : Proof.context -> string -> bool
       
    93   val is_prover_supported : Proof.context -> string -> bool
       
    94   val is_prover_installed : Proof.context -> string -> bool
       
    95   val remotify_prover_if_supported_and_not_already_remote :
       
    96     Proof.context -> string -> string option
       
    97   val remotify_prover_if_not_installed :
       
    98     Proof.context -> string -> string option
       
    99   val default_max_facts_of_prover : Proof.context -> string -> int
       
   100   val is_unit_equality : term -> bool
       
   101   val is_appropriate_prop_of_prover : Proof.context -> string -> term -> bool
       
   102   val weight_smt_fact :
       
   103     Proof.context -> int -> ((string * stature) * thm) * int
       
   104     -> (string * stature) * (int option * thm)
       
   105   val supported_provers : Proof.context -> unit
       
   106   val kill_provers : unit -> unit
       
   107   val running_provers : unit -> unit
       
   108   val messages : int option -> unit
       
   109   val is_fact_chained : (('a * stature) * 'b) -> bool
       
   110   val filter_used_facts :
       
   111     bool -> (''a * stature) list -> ((''a * stature) * 'b) list ->
       
   112     ((''a * stature) * 'b) list
       
   113   val isar_proof_reconstructor : Proof.context -> string -> string
       
   114   val get_prover : Proof.context -> mode -> string -> prover
       
   115 end;
       
   116 
       
   117 structure Sledgehammer_Prover : SLEDGEHAMMER_PROVER =
       
   118 struct
       
   119 
       
   120 open ATP_Util
       
   121 open ATP_Problem
       
   122 open ATP_Proof
       
   123 open ATP_Systems
       
   124 open ATP_Problem_Generate
       
   125 open ATP_Proof_Reconstruct
       
   126 open Metis_Tactic
       
   127 open Sledgehammer_Util
       
   128 open Sledgehammer_Fact
       
   129 open Sledgehammer_Reconstructor
       
   130 open Sledgehammer_Print
       
   131 open Sledgehammer_Reconstruct
       
   132 
       
   133 
       
   134 (** The Sledgehammer **)
       
   135 
       
   136 (* Empty string means create files in Isabelle's temporary files directory. *)
       
   137 val dest_dir = Attrib.setup_config_string @{binding sledgehammer_dest_dir} (K "")
       
   138 val problem_prefix = Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob")
       
   139 val completish = Attrib.setup_config_bool @{binding sledgehammer_completish} (K false)
       
   140 
       
   141 (* In addition to being easier to read, readable names are often much shorter,
       
   142    especially if types are mangled in names. This makes a difference for some
       
   143    provers (e.g., E). For these reason, short names are enabled by default. *)
       
   144 val atp_full_names = Attrib.setup_config_bool @{binding sledgehammer_atp_full_names} (K false)
       
   145 
       
   146 val smt_builtins = Attrib.setup_config_bool @{binding sledgehammer_smt_builtins} (K true)
       
   147 val smt_triggers = Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
       
   148 val smt_weights = Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true)
       
   149 val smt_weight_min_facts =
       
   150   Attrib.setup_config_int @{binding sledgehammer_smt_weight_min_facts} (K 20)
       
   151 
       
   152 datatype mode = Auto_Try | Try | Normal | MaSh | Auto_Minimize | Minimize
       
   153 
       
   154 (* Identifier that distinguishes Sledgehammer from other tools that could use
       
   155    "Async_Manager". *)
       
   156 val SledgehammerN = "Sledgehammer"
       
   157 
       
   158 val reconstructor_names = [metisN, smtN]
       
   159 val plain_metis = Metis (hd partial_type_encs, combsN)
       
   160 val is_reconstructor = member (op =) reconstructor_names
       
   161 
       
   162 val is_atp = member (op =) o supported_atps
       
   163 
       
   164 val select_smt_solver = Context.proof_map o SMT_Config.select_solver
       
   165 
       
   166 fun is_smt_prover ctxt = member (op =) (SMT_Solver.available_solvers_of ctxt)
       
   167 
       
   168 fun is_atp_of_format is_format ctxt name =
       
   169   let val thy = Proof_Context.theory_of ctxt in
       
   170     case try (get_atp thy) name of
       
   171       SOME config =>
       
   172       exists (fn (_, ((_, format, _, _, _), _)) => is_format format)
       
   173              (#best_slices (config ()) ctxt)
       
   174     | NONE => false
       
   175   end
       
   176 
       
   177 val is_unit_equational_atp = is_atp_of_format (curry (op =) CNF_UEQ)
       
   178 val is_ho_atp = is_atp_of_format is_format_higher_order
       
   179 
       
   180 fun is_prover_supported ctxt =
       
   181   let val thy = Proof_Context.theory_of ctxt in
       
   182     is_reconstructor orf is_atp thy orf is_smt_prover ctxt
       
   183   end
       
   184 
       
   185 fun is_prover_installed ctxt =
       
   186   is_reconstructor orf is_smt_prover ctxt orf
       
   187   is_atp_installed (Proof_Context.theory_of ctxt)
       
   188 
       
   189 fun remotify_prover_if_supported_and_not_already_remote ctxt name =
       
   190   if String.isPrefix remote_prefix name then
       
   191     SOME name
       
   192   else
       
   193     let val remote_name = remote_prefix ^ name in
       
   194       if is_prover_supported ctxt remote_name then SOME remote_name else NONE
       
   195     end
       
   196 
       
   197 fun remotify_prover_if_not_installed ctxt name =
       
   198   if is_prover_supported ctxt name andalso is_prover_installed ctxt name then
       
   199     SOME name
       
   200   else
       
   201     remotify_prover_if_supported_and_not_already_remote ctxt name
       
   202 
       
   203 fun get_slices slice slices =
       
   204   (0 upto length slices - 1) ~~ slices |> not slice ? (List.last #> single)
       
   205 
       
   206 val reconstructor_default_max_facts = 20
       
   207 
       
   208 fun slice_max_facts (_, ( ((max_facts, _), _, _, _, _), _)) = max_facts
       
   209 
       
   210 fun default_max_facts_of_prover ctxt name =
       
   211   let val thy = Proof_Context.theory_of ctxt in
       
   212     if is_reconstructor name then
       
   213       reconstructor_default_max_facts
       
   214     else if is_atp thy name then
       
   215       fold (Integer.max o slice_max_facts) (#best_slices (get_atp thy name ()) ctxt) 0
       
   216     else (* is_smt_prover ctxt name *)
       
   217       SMT_Solver.default_max_relevant ctxt name
       
   218   end
       
   219 
       
   220 fun is_if (@{const_name If}, _) = true
       
   221   | is_if _ = false
       
   222 
       
   223 (* Beware of "if and only if" (which is translated as such) and "If" (which is
       
   224    translated to conditional equations). *)
       
   225 fun is_good_unit_equality T t u =
       
   226   T <> @{typ bool} andalso not (exists (exists_Const is_if) [t, u])
       
   227 
       
   228 fun is_unit_equality (@{const Trueprop} $ t) = is_unit_equality t
       
   229   | is_unit_equality (Const (@{const_name all}, _) $ Abs (_, _, t)) =
       
   230     is_unit_equality t
       
   231   | is_unit_equality (Const (@{const_name All}, _) $ Abs (_, _, t)) =
       
   232     is_unit_equality t
       
   233   | is_unit_equality (Const (@{const_name "=="}, Type (_, [T, _])) $ t $ u) =
       
   234     is_good_unit_equality T t u
       
   235   | is_unit_equality (Const (@{const_name HOL.eq}, Type (_ , [T, _])) $ t $ u) =
       
   236     is_good_unit_equality T t u
       
   237   | is_unit_equality _ = false
       
   238 
       
   239 fun is_appropriate_prop_of_prover ctxt name =
       
   240   if is_unit_equational_atp ctxt name then is_unit_equality else K true
       
   241 
       
   242 fun supported_provers ctxt =
       
   243   let
       
   244     val thy = Proof_Context.theory_of ctxt
       
   245     val (remote_provers, local_provers) =
       
   246       reconstructor_names @
       
   247       sort_strings (supported_atps thy) @
       
   248       sort_strings (SMT_Solver.available_solvers_of ctxt)
       
   249       |> List.partition (String.isPrefix remote_prefix)
       
   250   in
       
   251     Output.urgent_message ("Supported provers: " ^
       
   252                            commas (local_provers @ remote_provers) ^ ".")
       
   253   end
       
   254 
       
   255 fun kill_provers () = Async_Manager.kill_threads SledgehammerN "prover"
       
   256 fun running_provers () = Async_Manager.running_threads SledgehammerN "prover"
       
   257 val messages = Async_Manager.thread_messages SledgehammerN "prover"
       
   258 
       
   259 
       
   260 (** problems, results, ATPs, etc. **)
       
   261 
       
   262 type params =
       
   263   {debug : bool,
       
   264    verbose : bool,
       
   265    overlord : bool,
       
   266    spy : bool,
       
   267    blocking : bool,
       
   268    provers : string list,
       
   269    type_enc : string option,
       
   270    strict : bool,
       
   271    lam_trans : string option,
       
   272    uncurried_aliases : bool option,
       
   273    learn : bool,
       
   274    fact_filter : string option,
       
   275    max_facts : int option,
       
   276    fact_thresholds : real * real,
       
   277    max_mono_iters : int option,
       
   278    max_new_mono_instances : int option,
       
   279    isar_proofs : bool option,
       
   280    compress_isar : real,
       
   281    try0_isar : bool,
       
   282    slice : bool,
       
   283    minimize : bool option,
       
   284    timeout : Time.time,
       
   285    preplay_timeout : Time.time,
       
   286    expect : string}
       
   287 
       
   288 type prover_problem =
       
   289   {comment : string,
       
   290    state : Proof.state,
       
   291    goal : thm,
       
   292    subgoal : int,
       
   293    subgoal_count : int,
       
   294    factss : (string * fact list) list}
       
   295 
       
   296 type prover_result =
       
   297   {outcome : atp_failure option,
       
   298    used_facts : (string * stature) list,
       
   299    used_from : fact list,
       
   300    run_time : Time.time,
       
   301    preplay : (reconstructor * play_outcome) Lazy.lazy,
       
   302    message : reconstructor * play_outcome -> string,
       
   303    message_tail : string}
       
   304 
       
   305 type prover =
       
   306   params -> ((string * string list) list -> string -> minimize_command)
       
   307   -> prover_problem -> prover_result
       
   308 
       
   309 (* FUDGE *)
       
   310 val smt_min_weight =
       
   311   Attrib.setup_config_int @{binding sledgehammer_smt_min_weight} (K 0)
       
   312 val smt_max_weight =
       
   313   Attrib.setup_config_int @{binding sledgehammer_smt_max_weight} (K 10)
       
   314 val smt_max_weight_index =
       
   315   Attrib.setup_config_int @{binding sledgehammer_smt_max_weight_index} (K 200)
       
   316 val smt_weight_curve = Unsynchronized.ref (fn x : int => x * x)
       
   317 
       
   318 fun smt_fact_weight ctxt j num_facts =
       
   319   if Config.get ctxt smt_weights andalso
       
   320      num_facts >= Config.get ctxt smt_weight_min_facts then
       
   321     let
       
   322       val min = Config.get ctxt smt_min_weight
       
   323       val max = Config.get ctxt smt_max_weight
       
   324       val max_index = Config.get ctxt smt_max_weight_index
       
   325       val curve = !smt_weight_curve
       
   326     in
       
   327       SOME (max - (max - min + 1) * curve (Int.max (0, max_index - j - 1))
       
   328             div curve max_index)
       
   329     end
       
   330   else
       
   331     NONE
       
   332 
       
   333 fun weight_smt_fact ctxt num_facts ((info, th), j) =
       
   334   let val thy = Proof_Context.theory_of ctxt in
       
   335     (info, (smt_fact_weight ctxt j num_facts, th |> Thm.transfer thy))
       
   336   end
       
   337 
       
   338 fun overlord_file_location_of_prover prover =
       
   339   (getenv "ISABELLE_HOME_USER", "prob_" ^ prover)
       
   340 
       
   341 fun proof_banner mode name =
       
   342   case mode of
       
   343     Auto_Try => "Auto Sledgehammer (" ^ quote name ^ ") found a proof"
       
   344   | Try => "Sledgehammer (" ^ quote name ^ ") found a proof"
       
   345   | _ => "Try this"
       
   346 
       
   347 fun bunch_of_reconstructors needs_full_types lam_trans =
       
   348   if needs_full_types then
       
   349     [Metis (full_type_enc, lam_trans false),
       
   350      Metis (really_full_type_enc, lam_trans false),
       
   351      Metis (full_type_enc, lam_trans true),
       
   352      Metis (really_full_type_enc, lam_trans true),
       
   353      SMT]
       
   354   else
       
   355     [Metis (partial_type_enc, lam_trans false),
       
   356      Metis (full_type_enc, lam_trans false),
       
   357      Metis (no_typesN, lam_trans true),
       
   358      Metis (really_full_type_enc, lam_trans true),
       
   359      SMT]
       
   360 
       
   361 fun extract_reconstructor ({type_enc, lam_trans, ...} : params) (Metis (type_enc', lam_trans')) =
       
   362     let
       
   363       val override_params =
       
   364         (if is_none type_enc andalso type_enc' = hd partial_type_encs then []
       
   365          else [("type_enc", [hd (unalias_type_enc type_enc')])]) @
       
   366         (if is_none lam_trans andalso lam_trans' = default_metis_lam_trans then []
       
   367          else [("lam_trans", [lam_trans'])])
       
   368     in (metisN, override_params) end
       
   369   | extract_reconstructor _ SMT = (smtN, [])
       
   370 
       
   371 (* based on "Mirabelle.can_apply" and generalized *)
       
   372 fun timed_apply timeout tac state i =
       
   373   let
       
   374     val {context = ctxt, facts, goal} = Proof.goal state
       
   375     val full_tac = Method.insert_tac facts i THEN tac ctxt i
       
   376   in
       
   377     TimeLimit.timeLimit timeout (try (Seq.pull o full_tac)) goal
       
   378   end
       
   379 
       
   380 fun tac_of_reconstructor (Metis (type_enc, lam_trans)) = metis_tac [type_enc] lam_trans
       
   381   | tac_of_reconstructor SMT = SMT_Solver.smt_tac
       
   382 
       
   383 fun timed_reconstructor reconstr debug timeout ths =
       
   384   timed_apply timeout (silence_reconstructors debug
       
   385     #> (fn ctxt => tac_of_reconstructor reconstr ctxt ths))
       
   386 
       
   387 fun is_fact_chained ((_, (sc, _)), _) = sc = Chained
       
   388 
       
   389 fun filter_used_facts keep_chained used =
       
   390   filter ((member (op =) used o fst) orf (if keep_chained then is_fact_chained else K false))
       
   391 
       
   392 fun play_one_line_proof mode debug verbose timeout pairs state i preferred reconstrs =
       
   393   let
       
   394     fun get_preferred reconstrs =
       
   395       if member (op =) reconstrs preferred then preferred
       
   396       else List.last reconstrs
       
   397   in
       
   398     if timeout = Time.zeroTime then
       
   399       (get_preferred reconstrs, Not_Played)
       
   400     else
       
   401       let
       
   402         val _ = if mode = Minimize then Output.urgent_message "Preplaying proof..." else ()
       
   403         val ths = pairs |> sort_wrt (fst o fst) |> map snd
       
   404         fun play [] [] = (get_preferred reconstrs, Play_Failed)
       
   405           | play timed_outs [] = (get_preferred timed_outs, Play_Timed_Out timeout)
       
   406           | play timed_out (reconstr :: reconstrs) =
       
   407             let
       
   408               val _ =
       
   409                 if verbose then
       
   410                   Output.urgent_message ("Trying \"" ^ string_of_reconstructor reconstr ^
       
   411                     "\" for " ^ string_of_time timeout ^ "...")
       
   412                 else
       
   413                   ()
       
   414               val timer = Timer.startRealTimer ()
       
   415             in
       
   416               case timed_reconstructor reconstr debug timeout ths state i of
       
   417                 SOME (SOME _) => (reconstr, Played (Timer.checkRealTimer timer))
       
   418               | _ => play timed_out reconstrs
       
   419             end
       
   420             handle TimeLimit.TimeOut => play (reconstr :: timed_out) reconstrs
       
   421       in
       
   422         play [] reconstrs
       
   423       end
       
   424   end
       
   425 
       
   426 
       
   427 (* generic TPTP-based ATPs *)
       
   428 
       
   429 (* Too general means, positive equality literal with a variable X as one
       
   430    operand, when X does not occur properly in the other operand. This rules out
       
   431    clearly inconsistent facts such as X = a | X = b, though it by no means
       
   432    guarantees soundness. *)
       
   433 
       
   434 fun get_facts_of_filter _ [(_, facts)] = facts
       
   435   | get_facts_of_filter fact_filter factss =
       
   436     case AList.lookup (op =) factss fact_filter of
       
   437       SOME facts => facts
       
   438     | NONE => snd (hd factss)
       
   439 
       
   440 (* Unwanted equalities are those between a (bound or schematic) variable that
       
   441    does not properly occur in the second operand. *)
       
   442 val is_exhaustive_finite =
       
   443   let
       
   444     fun is_bad_equal (Var z) t =
       
   445         not (exists_subterm (fn Var z' => z = z' | _ => false) t)
       
   446       | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j))
       
   447       | is_bad_equal _ _ = false
       
   448     fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1
       
   449     fun do_formula pos t =
       
   450       case (pos, t) of
       
   451         (_, @{const Trueprop} $ t1) => do_formula pos t1
       
   452       | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
       
   453         do_formula pos t'
       
   454       | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
       
   455         do_formula pos t'
       
   456       | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
       
   457         do_formula pos t'
       
   458       | (_, @{const "==>"} $ t1 $ t2) =>
       
   459         do_formula (not pos) t1 andalso
       
   460         (t2 = @{prop False} orelse do_formula pos t2)
       
   461       | (_, @{const HOL.implies} $ t1 $ t2) =>
       
   462         do_formula (not pos) t1 andalso
       
   463         (t2 = @{const False} orelse do_formula pos t2)
       
   464       | (_, @{const Not} $ t1) => do_formula (not pos) t1
       
   465       | (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
       
   466       | (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
       
   467       | (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2
       
   468       | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
       
   469       | _ => false
       
   470   in do_formula true end
       
   471 
       
   472 fun has_bound_or_var_of_type pred =
       
   473   exists_subterm (fn Var (_, T as Type _) => pred T
       
   474                    | Abs (_, T as Type _, _) => pred T
       
   475                    | _ => false)
       
   476 
       
   477 (* Facts are forbidden to contain variables of these types. The typical reason
       
   478    is that they lead to unsoundness. Note that "unit" satisfies numerous
       
   479    equations like "?x = ()". The resulting clauses will have no type constraint,
       
   480    yielding false proofs. Even "bool" leads to many unsound proofs, though only
       
   481    for higher-order problems. *)
       
   482 
       
   483 (* Facts containing variables of type "unit" or "bool" or of the form
       
   484    "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
       
   485    are omitted. *)
       
   486 fun is_dangerous_prop ctxt =
       
   487   transform_elim_prop
       
   488   #> (has_bound_or_var_of_type (is_type_surely_finite ctxt) orf
       
   489       is_exhaustive_finite)
       
   490 
       
   491 (* Important messages are important but not so important that users want to see
       
   492    them each time. *)
       
   493 val atp_important_message_keep_quotient = 25
       
   494 
       
   495 fun choose_type_enc strictness best_type_enc format =
       
   496   the_default best_type_enc
       
   497   #> type_enc_of_string strictness
       
   498   #> adjust_type_enc format
       
   499 
       
   500 fun isar_proof_reconstructor ctxt name =
       
   501   let val thy = Proof_Context.theory_of ctxt in
       
   502     if is_atp thy name then name
       
   503     else remotify_prover_if_not_installed ctxt eN |> the_default name
       
   504   end
       
   505 
       
   506 (* See the analogous logic in the function "maybe_minimize" in
       
   507   "sledgehammer_minimize.ML". *)
       
   508 fun choose_minimize_command ctxt (params as {isar_proofs, ...}) minimize_command name preplay =
       
   509   let
       
   510     val maybe_isar_name = name |> isar_proofs = SOME true ? isar_proof_reconstructor ctxt
       
   511     val (min_name, override_params) =
       
   512       (case preplay of
       
   513         (reconstr, Played _) =>
       
   514         if isar_proofs = SOME true then (maybe_isar_name, [])
       
   515         else extract_reconstructor params reconstr
       
   516       | _ => (maybe_isar_name, []))
       
   517   in minimize_command override_params min_name end
       
   518 
       
   519 val max_fact_instances = 10 (* FUDGE *)
       
   520 
       
   521 fun repair_monomorph_context max_iters best_max_iters max_new_instances
       
   522                              best_max_new_instances =
       
   523   Config.put Monomorph.max_rounds (max_iters |> the_default best_max_iters)
       
   524   #> Config.put Monomorph.max_new_instances
       
   525          (max_new_instances |> the_default best_max_new_instances)
       
   526   #> Config.put Monomorph.max_thm_instances max_fact_instances
       
   527 
       
   528 fun suffix_of_mode Auto_Try = "_try"
       
   529   | suffix_of_mode Try = "_try"
       
   530   | suffix_of_mode Normal = ""
       
   531   | suffix_of_mode MaSh = ""
       
   532   | suffix_of_mode Auto_Minimize = "_min"
       
   533   | suffix_of_mode Minimize = "_min"
       
   534 
       
   535 (* Give the ATPs some slack before interrupting them the hard way. "z3_tptp" on
       
   536    Linux appears to be the only ATP that does not honor its time limit. *)
       
   537 val atp_timeout_slack = seconds 1.0
       
   538 
       
   539 val mono_max_privileged_facts = 10
       
   540 
       
   541 (* For low values of "max_facts", this fudge value ensures that most slices are
       
   542    invoked with a nontrivial amount of facts. *)
       
   543 val max_fact_factor_fudge = 5
       
   544 
       
   545 fun run_atp mode name
       
   546     ({exec, arguments, proof_delims, known_failures, prem_role, best_slices, best_max_mono_iters,
       
   547       best_max_new_mono_instances, ...} : atp_config)
       
   548     (params as {debug, verbose, overlord, type_enc, strict, lam_trans, uncurried_aliases,
       
   549        fact_filter, max_facts, max_mono_iters, max_new_mono_instances, isar_proofs, compress_isar,
       
   550        try0_isar, slice, timeout, preplay_timeout, ...})
       
   551     minimize_command
       
   552     ({comment, state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
       
   553   let
       
   554     val thy = Proof.theory_of state
       
   555     val ctxt = Proof.context_of state
       
   556     val atp_mode =
       
   557       if Config.get ctxt completish then Sledgehammer_Completish
       
   558       else Sledgehammer
       
   559     val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal ctxt
       
   560     val (dest_dir, problem_prefix) =
       
   561       if overlord then overlord_file_location_of_prover name
       
   562       else (Config.get ctxt dest_dir, Config.get ctxt problem_prefix)
       
   563     val problem_file_name =
       
   564       Path.basic (problem_prefix ^ (if overlord then "" else serial_string ()) ^
       
   565                   suffix_of_mode mode ^ "_" ^ string_of_int subgoal)
       
   566     val prob_path =
       
   567       if dest_dir = "" then
       
   568         File.tmp_path problem_file_name
       
   569       else if File.exists (Path.explode dest_dir) then
       
   570         Path.append (Path.explode dest_dir) problem_file_name
       
   571       else
       
   572         error ("No such directory: " ^ quote dest_dir ^ ".")
       
   573     val exec = exec ()
       
   574     val command0 =
       
   575       case find_first (fn var => getenv var <> "") (fst exec) of
       
   576         SOME var =>
       
   577         let
       
   578           val pref = getenv var ^ "/"
       
   579           val paths = map (Path.explode o prefix pref) (snd exec)
       
   580         in
       
   581           case find_first File.exists paths of
       
   582             SOME path => path
       
   583           | NONE => error ("Bad executable: " ^ Path.print (hd paths) ^ ".")
       
   584         end
       
   585       | NONE => error ("The environment variable " ^ quote (List.last (fst exec)) ^
       
   586                        " is not set.")
       
   587     fun split_time s =
       
   588       let
       
   589         val split = String.tokens (fn c => str c = "\n")
       
   590         val (output, t) =
       
   591           s |> split |> (try split_last #> the_default ([], "0"))
       
   592             |>> cat_lines
       
   593         fun as_num f = f >> (fst o read_int)
       
   594         val num = as_num (Scan.many1 Symbol.is_ascii_digit)
       
   595         val digit = Scan.one Symbol.is_ascii_digit
       
   596         val num3 = as_num (digit ::: digit ::: (digit >> single))
       
   597         val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b)
       
   598         val as_time =
       
   599           raw_explode #> Scan.read Symbol.stopper time #> the_default 0
       
   600       in (output, as_time t |> Time.fromMilliseconds) end
       
   601     fun run () =
       
   602       let
       
   603         (* If slicing is disabled, we expand the last slice to fill the entire
       
   604            time available. *)
       
   605         val all_slices = best_slices ctxt
       
   606         val actual_slices = get_slices slice all_slices
       
   607         fun max_facts_of_slices f slices = fold (Integer.max o slice_max_facts o f) slices 0
       
   608         val num_actual_slices = length actual_slices
       
   609         val max_fact_factor =
       
   610           Real.fromInt (case max_facts of
       
   611               NONE => max_facts_of_slices I all_slices
       
   612             | SOME max => max)
       
   613           / Real.fromInt (max_facts_of_slices snd actual_slices)
       
   614         fun monomorphize_facts facts =
       
   615           let
       
   616             val ctxt =
       
   617               ctxt
       
   618               |> repair_monomorph_context max_mono_iters
       
   619                      best_max_mono_iters max_new_mono_instances
       
   620                      best_max_new_mono_instances
       
   621             (* pseudo-theorem involving the same constants as the subgoal *)
       
   622             val subgoal_th =
       
   623               Logic.list_implies (hyp_ts, concl_t) |> Skip_Proof.make_thm thy
       
   624             val rths =
       
   625               facts |> chop mono_max_privileged_facts
       
   626                     |>> map (pair 1 o snd)
       
   627                     ||> map (pair 2 o snd)
       
   628                     |> op @
       
   629                     |> cons (0, subgoal_th)
       
   630           in
       
   631             Monomorph.monomorph atp_schematic_consts_of ctxt rths
       
   632             |> tl |> curry ListPair.zip (map fst facts)
       
   633             |> maps (fn (name, rths) =>
       
   634                         map (pair name o zero_var_indexes o snd) rths)
       
   635           end
       
   636         fun run_slice time_left (cache_key, cache_value)
       
   637                 (slice, (time_frac,
       
   638                      (key as ((best_max_facts, best_fact_filter), format,
       
   639                               best_type_enc, best_lam_trans,
       
   640                               best_uncurried_aliases),
       
   641                       extra))) =
       
   642           let
       
   643             val effective_fact_filter =
       
   644               fact_filter |> the_default best_fact_filter
       
   645             val facts = get_facts_of_filter effective_fact_filter factss
       
   646             val num_facts =
       
   647               Real.ceil (max_fact_factor * Real.fromInt best_max_facts) +
       
   648               max_fact_factor_fudge
       
   649               |> Integer.min (length facts)
       
   650             val strictness = if strict then Strict else Non_Strict
       
   651             val type_enc =
       
   652               type_enc |> choose_type_enc strictness best_type_enc format
       
   653             val sound = is_type_enc_sound type_enc
       
   654             val real_ms = Real.fromInt o Time.toMilliseconds
       
   655             val slice_timeout =
       
   656               (real_ms time_left
       
   657                |> (if slice < num_actual_slices - 1 then
       
   658                      curry Real.min (time_frac * real_ms timeout)
       
   659                    else
       
   660                      I))
       
   661               * 0.001
       
   662               |> seconds
       
   663             val generous_slice_timeout =
       
   664               if mode = MaSh then one_day else Time.+ (slice_timeout, atp_timeout_slack)
       
   665             val _ =
       
   666               if debug then
       
   667                 quote name ^ " slice #" ^ string_of_int (slice + 1) ^
       
   668                 " with " ^ string_of_int num_facts ^ " fact" ^
       
   669                 plural_s num_facts ^ " for " ^ string_of_time slice_timeout ^ "..."
       
   670                 |> Output.urgent_message
       
   671               else
       
   672                 ()
       
   673             val readable_names = not (Config.get ctxt atp_full_names)
       
   674             val lam_trans =
       
   675               case lam_trans of
       
   676                 SOME s => s
       
   677               | NONE => best_lam_trans
       
   678             val uncurried_aliases =
       
   679               case uncurried_aliases of
       
   680                 SOME b => b
       
   681               | NONE => best_uncurried_aliases
       
   682             val value as (atp_problem, _, fact_names, _, _) =
       
   683               if cache_key = SOME key then
       
   684                 cache_value
       
   685               else
       
   686                 facts
       
   687                 |> not sound
       
   688                    ? filter_out (is_dangerous_prop ctxt o prop_of o snd)
       
   689                 |> take num_facts
       
   690                 |> not (is_type_enc_polymorphic type_enc) ? monomorphize_facts
       
   691                 |> map (apsnd prop_of)
       
   692                 |> prepare_atp_problem ctxt format prem_role type_enc atp_mode
       
   693                                        lam_trans uncurried_aliases
       
   694                                        readable_names true hyp_ts concl_t
       
   695             fun sel_weights () = atp_problem_selection_weights atp_problem
       
   696             fun ord_info () = atp_problem_term_order_info atp_problem
       
   697             val ord = effective_term_order ctxt name
       
   698             val full_proof = isar_proofs |> the_default (mode = Minimize)
       
   699             val args =
       
   700               arguments ctxt full_proof extra slice_timeout (File.shell_path prob_path)
       
   701                 (ord, ord_info, sel_weights)
       
   702             val command =
       
   703               "(exec 2>&1; " ^ File.shell_path command0 ^ " " ^ args ^ " " ^ ")"
       
   704               |> enclose "TIMEFORMAT='%3R'; { time " " ; }"
       
   705             val _ =
       
   706               atp_problem
       
   707               |> lines_of_atp_problem format ord ord_info
       
   708               |> cons ("% " ^ command ^ "\n" ^ (if comment = "" then "" else "% " ^ comment ^ "\n"))
       
   709               |> File.write_list prob_path
       
   710             val ((output, run_time), (atp_proof, outcome)) =
       
   711               TimeLimit.timeLimit generous_slice_timeout Isabelle_System.bash_output command
       
   712               |>> (if overlord then prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n") else I)
       
   713               |> fst |> split_time
       
   714               |> (fn accum as (output, _) =>
       
   715                      (accum,
       
   716                       extract_tstplike_proof_and_outcome verbose proof_delims known_failures output
       
   717                       |>> atp_proof_of_tstplike_proof atp_problem
       
   718                       handle UNRECOGNIZED_ATP_PROOF () => ([], SOME ProofIncomplete)))
       
   719               handle TimeLimit.TimeOut => (("", slice_timeout), ([], SOME TimedOut))
       
   720             val outcome =
       
   721               (case outcome of
       
   722                 NONE =>
       
   723                 (case used_facts_in_unsound_atp_proof ctxt fact_names atp_proof
       
   724                       |> Option.map (sort string_ord) of
       
   725                    SOME facts =>
       
   726                    let val failure = UnsoundProof (is_type_enc_sound type_enc, facts) in
       
   727                      if debug then (warning (string_of_atp_failure failure); NONE) else SOME failure
       
   728                    end
       
   729                  | NONE => NONE)
       
   730               | _ => outcome)
       
   731           in
       
   732             ((SOME key, value), (output, run_time, facts, atp_proof, outcome))
       
   733           end
       
   734         val timer = Timer.startRealTimer ()
       
   735         fun maybe_run_slice slice
       
   736                 (result as (cache, (_, run_time0, _, _, SOME _))) =
       
   737             let
       
   738               val time_left = Time.- (timeout, Timer.checkRealTimer timer)
       
   739             in
       
   740               if Time.<= (time_left, Time.zeroTime) then
       
   741                 result
       
   742               else
       
   743                 run_slice time_left cache slice
       
   744                 |> (fn (cache, (output, run_time, used_from, atp_proof, outcome)) =>
       
   745                   (cache, (output, Time.+ (run_time0, run_time), used_from, atp_proof, outcome)))
       
   746             end
       
   747           | maybe_run_slice _ result = result
       
   748       in
       
   749         ((NONE, ([], Symtab.empty, Vector.fromList [], [], Symtab.empty)),
       
   750          ("", Time.zeroTime, [], [], SOME InternalError))
       
   751         |> fold maybe_run_slice actual_slices
       
   752       end
       
   753     (* If the problem file has not been exported, remove it; otherwise, export
       
   754        the proof file too. *)
       
   755     fun clean_up () =
       
   756       if dest_dir = "" then (try File.rm prob_path; ()) else ()
       
   757     fun export (_, (output, _, _, _, _)) =
       
   758       if dest_dir = "" then ()
       
   759       else File.write (Path.explode (Path.implode prob_path ^ "_proof")) output
       
   760     val ((_, (_, pool, fact_names, lifted, sym_tab)),
       
   761          (output, run_time, used_from, atp_proof, outcome)) =
       
   762       with_cleanup clean_up run () |> tap export
       
   763     val important_message =
       
   764       if mode = Normal andalso random_range 0 (atp_important_message_keep_quotient - 1) = 0
       
   765       then
       
   766         extract_important_message output
       
   767       else
       
   768         ""
       
   769     val (used_facts, preplay, message, message_tail) =
       
   770       (case outcome of
       
   771         NONE =>
       
   772         let
       
   773           val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof
       
   774           val needs_full_types = is_typed_helper_used_in_atp_proof atp_proof
       
   775           val reconstrs =
       
   776             bunch_of_reconstructors needs_full_types (lam_trans_of_atp_proof atp_proof
       
   777               o (fn desperate => if desperate then hide_lamsN else default_metis_lam_trans))
       
   778         in
       
   779           (used_facts,
       
   780            Lazy.lazy (fn () =>
       
   781              let val used_pairs = used_from |> filter_used_facts false used_facts in
       
   782                play_one_line_proof mode debug verbose preplay_timeout used_pairs state subgoal
       
   783                  (hd reconstrs) reconstrs
       
   784              end),
       
   785            fn preplay =>
       
   786               let
       
   787                 val _ = if verbose then Output.urgent_message "Generating proof text..." else ()
       
   788                 fun isar_params () =
       
   789                   let
       
   790                     val metis_type_enc =
       
   791                       if is_typed_helper_used_in_atp_proof atp_proof then full_typesN
       
   792                       else partial_typesN
       
   793                     val metis_lam_trans = lam_trans_of_atp_proof atp_proof default_metis_lam_trans
       
   794                     val atp_proof =
       
   795                       atp_proof
       
   796                       |> termify_atp_proof ctxt pool lifted sym_tab
       
   797                       |> introduce_spass_skolem
       
   798                       |> factify_atp_proof fact_names hyp_ts concl_t
       
   799                   in
       
   800                     (verbose, metis_type_enc, metis_lam_trans, preplay_timeout, compress_isar,
       
   801                      try0_isar, atp_proof, goal)
       
   802                   end
       
   803                 val one_line_params =
       
   804                   (preplay, proof_banner mode name, used_facts,
       
   805                    choose_minimize_command ctxt params minimize_command name preplay,
       
   806                    subgoal, subgoal_count)
       
   807                 val num_chained = length (#facts (Proof.goal state))
       
   808               in
       
   809                 proof_text ctxt debug isar_proofs isar_params num_chained one_line_params
       
   810               end,
       
   811            (if verbose then "\nATP real CPU time: " ^ string_of_time run_time ^ "." else "") ^
       
   812            (if important_message <> "" then
       
   813               "\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^ important_message
       
   814             else
       
   815               ""))
       
   816         end
       
   817       | SOME failure =>
       
   818         ([], Lazy.value (plain_metis, Play_Failed), fn _ => string_of_atp_failure failure, ""))
       
   819   in
       
   820     {outcome = outcome, used_facts = used_facts, used_from = used_from, run_time = run_time,
       
   821      preplay = preplay, message = message, message_tail = message_tail}
       
   822   end
       
   823 
       
   824 (* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until
       
   825    these are sorted out properly in the SMT module, we have to interpret these
       
   826    ourselves. *)
       
   827 val remote_smt_failures =
       
   828   [(2, NoLibwwwPerl),
       
   829    (22, CantConnect)]
       
   830 val z3_failures =
       
   831   [(101, OutOfResources),
       
   832    (103, MalformedInput),
       
   833    (110, MalformedInput),
       
   834    (112, TimedOut)]
       
   835 val unix_failures =
       
   836   [(138, Crashed),
       
   837    (139, Crashed)]
       
   838 val smt_failures = remote_smt_failures @ z3_failures @ unix_failures
       
   839 
       
   840 fun failure_of_smt_failure (SMT_Failure.Counterexample {is_real_cex, ...}) =
       
   841     if is_real_cex then Unprovable else GaveUp
       
   842   | failure_of_smt_failure SMT_Failure.Time_Out = TimedOut
       
   843   | failure_of_smt_failure (SMT_Failure.Abnormal_Termination code) =
       
   844     (case AList.lookup (op =) smt_failures code of
       
   845        SOME failure => failure
       
   846      | NONE => UnknownError ("Abnormal termination with exit code " ^
       
   847                              string_of_int code ^ "."))
       
   848   | failure_of_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources
       
   849   | failure_of_smt_failure (SMT_Failure.Other_Failure s) = UnknownError s
       
   850 
       
   851 (* FUDGE *)
       
   852 val smt_max_slices =
       
   853   Attrib.setup_config_int @{binding sledgehammer_smt_max_slices} (K 8)
       
   854 val smt_slice_fact_frac =
       
   855   Attrib.setup_config_real @{binding sledgehammer_smt_slice_fact_frac}
       
   856                            (K 0.667)
       
   857 val smt_slice_time_frac =
       
   858   Attrib.setup_config_real @{binding sledgehammer_smt_slice_time_frac} (K 0.333)
       
   859 val smt_slice_min_secs =
       
   860   Attrib.setup_config_int @{binding sledgehammer_smt_slice_min_secs} (K 3)
       
   861 
       
   862 val is_boring_builtin_typ =
       
   863   not o exists_subtype (member (op =) [@{typ nat}, @{typ int}, HOLogic.realT])
       
   864 
       
   865 fun smt_filter_loop name ({debug, overlord, max_mono_iters, max_new_mono_instances, timeout, slice,
       
   866       ...} : params) state goal i =
       
   867   let
       
   868     fun repair_context ctxt =
       
   869       ctxt |> select_smt_solver name
       
   870            |> Config.put SMT_Config.verbose debug
       
   871            |> (if overlord then
       
   872                  Config.put SMT_Config.debug_files
       
   873                             (overlord_file_location_of_prover name
       
   874                              |> (fn (path, name) => path ^ "/" ^ name))
       
   875                else
       
   876                  I)
       
   877            |> Config.put SMT_Config.infer_triggers
       
   878                          (Config.get ctxt smt_triggers)
       
   879            |> not (Config.get ctxt smt_builtins)
       
   880               ? (SMT_Builtin.filter_builtins is_boring_builtin_typ
       
   881                  #> Config.put SMT_Config.datatypes false)
       
   882            |> repair_monomorph_context max_mono_iters default_max_mono_iters
       
   883                   max_new_mono_instances default_max_new_mono_instances
       
   884     val state = Proof.map_context (repair_context) state
       
   885     val ctxt = Proof.context_of state
       
   886     val max_slices = if slice then Config.get ctxt smt_max_slices else 1
       
   887     fun do_slice timeout slice outcome0 time_so_far
       
   888                  (weighted_factss as (fact_filter, weighted_facts) :: _) =
       
   889       let
       
   890         val timer = Timer.startRealTimer ()
       
   891         val slice_timeout =
       
   892           if slice < max_slices then
       
   893             let val ms = Time.toMilliseconds timeout in
       
   894               Int.min (ms,
       
   895                   Int.max (1000 * Config.get ctxt smt_slice_min_secs,
       
   896                       Real.ceil (Config.get ctxt smt_slice_time_frac
       
   897                                  * Real.fromInt ms)))
       
   898               |> Time.fromMilliseconds
       
   899             end
       
   900           else
       
   901             timeout
       
   902         val num_facts = length weighted_facts
       
   903         val _ =
       
   904           if debug then
       
   905             quote name ^ " slice " ^ string_of_int slice ^ " with " ^ string_of_int num_facts ^
       
   906             " fact" ^ plural_s num_facts ^ " for " ^ string_of_time slice_timeout
       
   907             |> Output.urgent_message
       
   908           else
       
   909             ()
       
   910         val birth = Timer.checkRealTimer timer
       
   911         val _ =
       
   912           if debug then Output.urgent_message "Invoking SMT solver..." else ()
       
   913         val (outcome, used_facts) =
       
   914           SMT_Solver.smt_filter_preprocess ctxt [] goal weighted_facts i
       
   915           |> SMT_Solver.smt_filter_apply slice_timeout
       
   916           |> (fn {outcome, used_facts} => (outcome, used_facts))
       
   917           handle exn => if Exn.is_interrupt exn then
       
   918                           reraise exn
       
   919                         else
       
   920                           (ML_Compiler.exn_message exn
       
   921                            |> SMT_Failure.Other_Failure |> SOME, [])
       
   922         val death = Timer.checkRealTimer timer
       
   923         val outcome0 = if is_none outcome0 then SOME outcome else outcome0
       
   924         val time_so_far = Time.+ (time_so_far, Time.- (death, birth))
       
   925         val too_many_facts_perhaps =
       
   926           case outcome of
       
   927             NONE => false
       
   928           | SOME (SMT_Failure.Counterexample _) => false
       
   929           | SOME SMT_Failure.Time_Out => slice_timeout <> timeout
       
   930           | SOME (SMT_Failure.Abnormal_Termination _) => true (* kind of *)
       
   931           | SOME SMT_Failure.Out_Of_Memory => true
       
   932           | SOME (SMT_Failure.Other_Failure _) => true
       
   933         val timeout = Time.- (timeout, Timer.checkRealTimer timer)
       
   934       in
       
   935         if too_many_facts_perhaps andalso slice < max_slices andalso num_facts > 0 andalso
       
   936            Time.> (timeout, Time.zeroTime) then
       
   937           let
       
   938             val new_num_facts =
       
   939               Real.ceil (Config.get ctxt smt_slice_fact_frac * Real.fromInt num_facts)
       
   940             val weighted_factss as (new_fact_filter, _) :: _ =
       
   941               weighted_factss
       
   942               |> (fn (x :: xs) => xs @ [x])
       
   943               |> app_hd (apsnd (take new_num_facts))
       
   944             val show_filter = fact_filter <> new_fact_filter
       
   945             fun num_of_facts fact_filter num_facts =
       
   946               string_of_int num_facts ^
       
   947               (if show_filter then " " ^ quote fact_filter else "") ^
       
   948               " fact" ^ plural_s num_facts
       
   949             val _ =
       
   950               if debug then
       
   951                 quote name ^ " invoked with " ^
       
   952                 num_of_facts fact_filter num_facts ^ ": " ^
       
   953                 string_of_atp_failure (failure_of_smt_failure (the outcome)) ^
       
   954                 " Retrying with " ^ num_of_facts new_fact_filter new_num_facts ^
       
   955                 "..."
       
   956                 |> Output.urgent_message
       
   957               else
       
   958                 ()
       
   959           in
       
   960             do_slice timeout (slice + 1) outcome0 time_so_far weighted_factss
       
   961           end
       
   962         else
       
   963           {outcome = if is_none outcome then NONE else the outcome0, used_facts = used_facts,
       
   964            used_from = map (apsnd snd) weighted_facts, run_time = time_so_far}
       
   965       end
       
   966   in
       
   967     do_slice timeout 1 NONE Time.zeroTime
       
   968   end
       
   969 
       
   970 fun run_smt_solver mode name (params as {debug, verbose, preplay_timeout, ...}) minimize_command
       
   971     ({state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
       
   972   let
       
   973     val ctxt = Proof.context_of state
       
   974     fun weight_facts facts =
       
   975       let val num_facts = length facts in
       
   976         facts ~~ (0 upto num_facts - 1)
       
   977         |> map (weight_smt_fact ctxt num_facts)
       
   978       end
       
   979     val weighted_factss = factss |> map (apsnd weight_facts)
       
   980     val {outcome, used_facts = used_pairs, used_from, run_time} =
       
   981       smt_filter_loop name params state goal subgoal weighted_factss
       
   982     val used_facts = used_pairs |> map fst
       
   983     val outcome = outcome |> Option.map failure_of_smt_failure
       
   984     val (preplay, message, message_tail) =
       
   985       case outcome of
       
   986         NONE =>
       
   987         (Lazy.lazy (fn () =>
       
   988            play_one_line_proof mode debug verbose preplay_timeout used_pairs
       
   989                state subgoal SMT
       
   990                (bunch_of_reconstructors false (fn desperate =>
       
   991                   if desperate then liftingN else default_metis_lam_trans))),
       
   992          fn preplay =>
       
   993             let
       
   994               val one_line_params =
       
   995                 (preplay, proof_banner mode name, used_facts,
       
   996                  choose_minimize_command ctxt params minimize_command name preplay,
       
   997                  subgoal, subgoal_count)
       
   998               val num_chained = length (#facts (Proof.goal state))
       
   999             in
       
  1000               one_line_proof_text num_chained one_line_params
       
  1001             end,
       
  1002          if verbose then "\nSMT solver real CPU time: " ^ string_of_time run_time ^ "." else "")
       
  1003       | SOME failure =>
       
  1004         (Lazy.value (plain_metis, Play_Failed),
       
  1005          fn _ => string_of_atp_failure failure, "")
       
  1006   in
       
  1007     {outcome = outcome, used_facts = used_facts, used_from = used_from,
       
  1008      run_time = run_time, preplay = preplay, message = message,
       
  1009      message_tail = message_tail}
       
  1010   end
       
  1011 
       
  1012 fun run_reconstructor mode name
       
  1013         (params as {debug, verbose, timeout, type_enc, lam_trans, ...})
       
  1014         minimize_command
       
  1015         ({state, subgoal, subgoal_count, factss = (_, facts) :: _, ...}
       
  1016          : prover_problem) =
       
  1017   let
       
  1018     val reconstr =
       
  1019       if name = metisN then
       
  1020         Metis (type_enc |> the_default (hd partial_type_encs),
       
  1021                lam_trans |> the_default default_metis_lam_trans)
       
  1022       else if name = smtN then
       
  1023         SMT
       
  1024       else
       
  1025         raise Fail ("unknown reconstructor: " ^ quote name)
       
  1026     val used_facts = facts |> map fst
       
  1027   in
       
  1028     (case play_one_line_proof (if mode = Minimize then Normal else mode) debug verbose timeout facts
       
  1029        state subgoal reconstr [reconstr] of
       
  1030       play as (_, Played time) =>
       
  1031       {outcome = NONE, used_facts = used_facts, used_from = facts, run_time = time,
       
  1032        preplay = Lazy.value play,
       
  1033        message =
       
  1034          fn play =>
       
  1035             let
       
  1036               val (_, override_params) = extract_reconstructor params reconstr
       
  1037               val one_line_params =
       
  1038                 (play, proof_banner mode name, used_facts, minimize_command override_params name,
       
  1039                  subgoal, subgoal_count)
       
  1040               val num_chained = length (#facts (Proof.goal state))
       
  1041             in
       
  1042               one_line_proof_text num_chained one_line_params
       
  1043             end,
       
  1044        message_tail = ""}
       
  1045     | play =>
       
  1046       let
       
  1047         val failure = (case play of (_, Play_Failed) => GaveUp | _ => TimedOut)
       
  1048       in
       
  1049         {outcome = SOME failure, used_facts = [], used_from = [],
       
  1050          run_time = Time.zeroTime, preplay = Lazy.value play,
       
  1051          message = fn _ => string_of_atp_failure failure, message_tail = ""}
       
  1052       end)
       
  1053   end
       
  1054 
       
  1055 fun get_prover ctxt mode name =
       
  1056   let val thy = Proof_Context.theory_of ctxt in
       
  1057     if is_reconstructor name then run_reconstructor mode name
       
  1058     else if is_atp thy name then run_atp mode name (get_atp thy name ())
       
  1059     else if is_smt_prover ctxt name then run_smt_solver mode name
       
  1060     else error ("No such prover: " ^ name ^ ".")
       
  1061   end
       
  1062 
       
  1063 end;