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