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