src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML
author blanchet
Tue, 26 Jun 2012 11:14:39 +0200
changeset 49145 defbcdc60fd6
parent 49119 d2173ff80c57
child 49146 1016664b8feb
permissions -rw-r--r--
tuning
     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 option,
    34      max_new_mono_instances: int option,
    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 
    54   type prover_result =
    55     {outcome: failure option,
    56      used_facts: (string * stature) list,
    57      run_time: Time.time,
    58      preplay: unit -> play,
    59      message: play -> string,
    60      message_tail: string}
    61 
    62   type prover =
    63     params -> ((string * string list) list -> string -> minimize_command)
    64     -> prover_problem -> prover_result
    65 
    66   val dest_dir : string Config.T
    67   val problem_prefix : string Config.T
    68   val aggressive : bool 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 config =>
   151       exists (fn (_, (_, ((_, format, _, _, _), _))) => is_format format)
   152              (#best_slices (config ()) 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.05,
   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 option,
   299    max_new_mono_instances: int option,
   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 
   319 type prover_result =
   320   {outcome: failure option,
   321    used_facts: (string * stature) list,
   322    run_time: Time.time,
   323    preplay: unit -> play,
   324    message: play -> string,
   325    message_tail: string}
   326 
   327 type prover =
   328   params -> ((string * string list) list -> string -> minimize_command)
   329   -> prover_problem -> prover_result
   330 
   331 (* configuration attributes *)
   332 
   333 (* Empty string means create files in Isabelle's temporary files directory. *)
   334 val dest_dir =
   335   Attrib.setup_config_string @{binding sledgehammer_dest_dir} (K "")
   336 val problem_prefix =
   337   Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob")
   338 val aggressive =
   339   Attrib.setup_config_bool @{binding sledgehammer_aggressive} (K false)
   340 
   341 (* In addition to being easier to read, readable names are often much shorter,
   342    especially if types are mangled in names. This makes a difference for some
   343    provers (e.g., E). For these reason, short names are enabled by default. *)
   344 val atp_full_names =
   345   Attrib.setup_config_bool @{binding sledgehammer_atp_full_names} (K false)
   346 
   347 val smt_triggers =
   348   Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
   349 val smt_weights =
   350   Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true)
   351 val smt_weight_min_facts =
   352   Attrib.setup_config_int @{binding sledgehammer_smt_weight_min_facts} (K 20)
   353 
   354 (* FUDGE *)
   355 val smt_min_weight =
   356   Attrib.setup_config_int @{binding sledgehammer_smt_min_weight} (K 0)
   357 val smt_max_weight =
   358   Attrib.setup_config_int @{binding sledgehammer_smt_max_weight} (K 10)
   359 val smt_max_weight_index =
   360   Attrib.setup_config_int @{binding sledgehammer_smt_max_weight_index} (K 200)
   361 val smt_weight_curve = Unsynchronized.ref (fn x : int => x * x)
   362 
   363 fun smt_fact_weight ctxt j num_facts =
   364   if Config.get ctxt smt_weights andalso
   365      num_facts >= Config.get ctxt smt_weight_min_facts then
   366     let
   367       val min = Config.get ctxt smt_min_weight
   368       val max = Config.get ctxt smt_max_weight
   369       val max_index = Config.get ctxt smt_max_weight_index
   370       val curve = !smt_weight_curve
   371     in
   372       SOME (max - (max - min + 1) * curve (Int.max (0, max_index - j - 1))
   373             div curve max_index)
   374     end
   375   else
   376     NONE
   377 
   378 fun weight_smt_fact ctxt num_facts ((info, th), j) =
   379   let val thy = Proof_Context.theory_of ctxt in
   380     (info, (smt_fact_weight ctxt j num_facts, th |> Thm.transfer thy))
   381   end
   382 
   383 fun untranslated_fact (Untranslated_Fact p) = p
   384   | untranslated_fact (SMT_Weighted_Fact (info, (_, th))) = (info, th)
   385 fun smt_weighted_fact _ _ (SMT_Weighted_Fact p, _) = p
   386   | smt_weighted_fact ctxt num_facts (fact, j) =
   387     (untranslated_fact fact, j) |> weight_smt_fact ctxt num_facts
   388 
   389 fun overlord_file_location_for_prover prover =
   390   (getenv "ISABELLE_HOME_USER", "prob_" ^ prover)
   391 
   392 fun with_path cleanup after f path =
   393   Exn.capture f path
   394   |> tap (fn _ => cleanup path)
   395   |> Exn.release
   396   |> tap (after path)
   397 
   398 fun proof_banner mode name =
   399   case mode of
   400     Auto_Try => "Auto Sledgehammer (" ^ quote name ^ ") found a proof"
   401   | Try => "Sledgehammer (" ^ quote name ^ ") found a proof"
   402   | _ => "Try this"
   403 
   404 fun bunch_of_reconstructors needs_full_types lam_trans =
   405   [(false, Metis (partial_type_enc, lam_trans false)),
   406    (true, Metis (full_type_enc, lam_trans false)),
   407    (false, Metis (no_typesN, lam_trans true)),
   408    (true, Metis (really_full_type_enc, lam_trans true)),
   409    (true, SMT)]
   410   |> map_filter (fn (full_types, reconstr) =>
   411                     if needs_full_types andalso not full_types then NONE
   412                     else SOME reconstr)
   413 
   414 fun extract_reconstructor ({type_enc, lam_trans, ...} : params)
   415                           (Metis (type_enc', lam_trans')) =
   416     let
   417       val override_params =
   418         (if is_none type_enc andalso type_enc' = hd partial_type_encs then
   419            []
   420          else
   421            [("type_enc", [hd (unalias_type_enc type_enc')])]) @
   422         (if is_none lam_trans andalso lam_trans' = metis_default_lam_trans then
   423            []
   424          else
   425            [("lam_trans", [lam_trans'])])
   426     in (metisN, override_params) end
   427   | extract_reconstructor _ SMT = (smtN, [])
   428 
   429 (* based on "Mirabelle.can_apply" and generalized *)
   430 fun timed_apply timeout tac state i =
   431   let
   432     val {context = ctxt, facts, goal} = Proof.goal state
   433     val full_tac = Method.insert_tac facts i THEN tac ctxt i
   434   in TimeLimit.timeLimit timeout (try (Seq.pull o full_tac)) goal end
   435 
   436 fun tac_for_reconstructor (Metis (type_enc, lam_trans)) =
   437     metis_tac [type_enc] lam_trans
   438   | tac_for_reconstructor SMT = SMT_Solver.smt_tac
   439 
   440 fun timed_reconstructor reconstr debug timeout ths =
   441   (Config.put Metis_Tactic.verbose debug
   442    #> Config.put SMT_Config.verbose debug
   443    #> (fn ctxt => tac_for_reconstructor reconstr ctxt ths))
   444   |> timed_apply timeout
   445 
   446 fun filter_used_facts used = filter (member (op =) used o fst)
   447 
   448 fun play_one_line_proof mode debug verbose timeout pairs state i preferred
   449                         reconstrs =
   450   let
   451     val _ =
   452       if mode = Minimize andalso Time.> (timeout, Time.zeroTime) then
   453         Output.urgent_message "Preplaying proof..."
   454       else
   455         ()
   456     val ths = pairs |> sort_wrt (fst o fst) |> map snd
   457     fun get_preferred reconstrs =
   458       if member (op =) reconstrs preferred then preferred
   459       else List.last reconstrs
   460     fun play [] [] = Failed_to_Play (get_preferred reconstrs)
   461       | play timed_outs [] =
   462         Trust_Playable (get_preferred timed_outs, SOME timeout)
   463       | play timed_out (reconstr :: reconstrs) =
   464         let
   465           val _ =
   466             if verbose then
   467               "Trying \"" ^ string_for_reconstructor reconstr ^ "\" for " ^
   468               string_from_time timeout ^ "..."
   469               |> Output.urgent_message
   470             else
   471               ()
   472           val timer = Timer.startRealTimer ()
   473         in
   474           case timed_reconstructor reconstr debug timeout ths state i of
   475             SOME (SOME _) => Played (reconstr, Timer.checkRealTimer timer)
   476           | _ => play timed_out reconstrs
   477         end
   478         handle TimeLimit.TimeOut => play (reconstr :: timed_out) reconstrs
   479   in
   480     if timeout = Time.zeroTime then
   481       Trust_Playable (get_preferred reconstrs, NONE)
   482     else
   483       play [] reconstrs
   484   end
   485 
   486 
   487 (* generic TPTP-based ATPs *)
   488 
   489 (* Too general means, positive equality literal with a variable X as one
   490    operand, when X does not occur properly in the other operand. This rules out
   491    clearly inconsistent facts such as X = a | X = b, though it by no means
   492    guarantees soundness. *)
   493 
   494 (* Unwanted equalities are those between a (bound or schematic) variable that
   495    does not properly occur in the second operand. *)
   496 val is_exhaustive_finite =
   497   let
   498     fun is_bad_equal (Var z) t =
   499         not (exists_subterm (fn Var z' => z = z' | _ => false) t)
   500       | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j))
   501       | is_bad_equal _ _ = false
   502     fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1
   503     fun do_formula pos t =
   504       case (pos, t) of
   505         (_, @{const Trueprop} $ t1) => do_formula pos t1
   506       | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
   507         do_formula pos t'
   508       | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
   509         do_formula pos t'
   510       | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
   511         do_formula pos t'
   512       | (_, @{const "==>"} $ t1 $ t2) =>
   513         do_formula (not pos) t1 andalso
   514         (t2 = @{prop False} orelse do_formula pos t2)
   515       | (_, @{const HOL.implies} $ t1 $ t2) =>
   516         do_formula (not pos) t1 andalso
   517         (t2 = @{const False} orelse do_formula pos t2)
   518       | (_, @{const Not} $ t1) => do_formula (not pos) t1
   519       | (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
   520       | (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
   521       | (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2
   522       | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
   523       | _ => false
   524   in do_formula true end
   525 
   526 fun has_bound_or_var_of_type pred =
   527   exists_subterm (fn Var (_, T as Type _) => pred T
   528                    | Abs (_, T as Type _, _) => pred T
   529                    | _ => false)
   530 
   531 (* Facts are forbidden to contain variables of these types. The typical reason
   532    is that they lead to unsoundness. Note that "unit" satisfies numerous
   533    equations like "?x = ()". The resulting clauses will have no type constraint,
   534    yielding false proofs. Even "bool" leads to many unsound proofs, though only
   535    for higher-order problems. *)
   536 
   537 (* Facts containing variables of type "unit" or "bool" or of the form
   538    "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
   539    are omitted. *)
   540 fun is_dangerous_prop ctxt =
   541   transform_elim_prop
   542   #> (has_bound_or_var_of_type (is_type_surely_finite ctxt) orf
   543       is_exhaustive_finite)
   544 
   545 (* Important messages are important but not so important that users want to see
   546    them each time. *)
   547 val atp_important_message_keep_quotient = 25
   548 
   549 fun choose_type_enc soundness best_type_enc format =
   550   the_default best_type_enc
   551   #> type_enc_from_string soundness
   552   #> adjust_type_enc format
   553 
   554 val metis_minimize_max_time = seconds 2.0
   555 
   556 fun choose_minimize_command params minimize_command name preplay =
   557   let
   558     val (name, override_params) =
   559       case preplay of
   560         Played (reconstr, time) =>
   561         if Time.<= (time, metis_minimize_max_time) then
   562           extract_reconstructor params reconstr
   563         else
   564           (name, [])
   565       | _ => (name, [])
   566   in minimize_command override_params name end
   567 
   568 fun repair_monomorph_context max_iters best_max_iters max_new_instances
   569                              best_max_new_instances =
   570   Config.put Monomorph.max_rounds (max_iters |> the_default best_max_iters)
   571   #> Config.put Monomorph.max_new_instances
   572          (max_new_instances |> the_default best_max_new_instances)
   573   #> Config.put Monomorph.keep_partial_instances false
   574 
   575 fun suffix_for_mode Auto_Try = "_auto_try"
   576   | suffix_for_mode Try = "_try"
   577   | suffix_for_mode Normal = ""
   578   | suffix_for_mode Auto_Minimize = "_auto_min"
   579   | suffix_for_mode Minimize = "_min"
   580 
   581 (* Give the ATPs some slack before interrupting them the hard way. "z3_tptp" on
   582    Linux appears to be the only ATP that does not honor its time limit. *)
   583 val atp_timeout_slack = seconds 1.0
   584 
   585 fun run_atp mode name
   586         ({exec, required_vars, arguments, proof_delims, known_failures,
   587           prem_role, best_slices, best_max_mono_iters,
   588           best_max_new_mono_instances, ...} : atp_config)
   589         (params as {debug, verbose, overlord, type_enc, strict, lam_trans,
   590                     uncurried_aliases, max_relevant, max_mono_iters,
   591                     max_new_mono_instances, isar_proof, isar_shrink_factor,
   592                     slice, timeout, preplay_timeout, ...})
   593         minimize_command
   594         ({state, goal, subgoal, subgoal_count, facts, ...} : prover_problem) =
   595   let
   596     val thy = Proof.theory_of state
   597     val ctxt = Proof.context_of state
   598     val atp_mode =
   599       if Config.get ctxt aggressive then Sledgehammer_Aggressive
   600       else Sledgehammer
   601     val (_, hyp_ts, concl_t) = strip_subgoal ctxt goal subgoal
   602     val (dest_dir, problem_prefix) =
   603       if overlord then overlord_file_location_for_prover name
   604       else (Config.get ctxt dest_dir, Config.get ctxt problem_prefix)
   605     val problem_file_name =
   606       Path.basic (problem_prefix ^ (if overlord then "" else serial_string ()) ^
   607                   suffix_for_mode mode ^ "_" ^ string_of_int subgoal)
   608     val problem_path_name =
   609       if dest_dir = "" then
   610         File.tmp_path problem_file_name
   611       else if File.exists (Path.explode dest_dir) then
   612         Path.append (Path.explode dest_dir) problem_file_name
   613       else
   614         error ("No such directory: " ^ quote dest_dir ^ ".")
   615     val command =
   616       case find_first (fn var => getenv var <> "") (fst exec) of
   617         SOME var => Path.explode (getenv var ^ "/" ^ snd exec)
   618       | NONE => error ("The environment variable " ^ quote (hd (fst exec)) ^
   619                        " is not set.")
   620     fun split_time s =
   621       let
   622         val split = String.tokens (fn c => str c = "\n")
   623         val (output, t) =
   624           s |> split |> (try split_last #> the_default ([], "0"))
   625             |>> cat_lines
   626         fun as_num f = f >> (fst o read_int)
   627         val num = as_num (Scan.many1 Symbol.is_ascii_digit)
   628         val digit = Scan.one Symbol.is_ascii_digit
   629         val num3 = as_num (digit ::: digit ::: (digit >> single))
   630         val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b)
   631         val as_time =
   632           raw_explode #> Scan.read Symbol.stopper time #> the_default 0
   633       in (output, as_time t |> Time.fromMilliseconds) end
   634     fun run_on prob_file =
   635       case find_first (forall (fn var => getenv var = ""))
   636                       (fst exec :: required_vars) of
   637         SOME home_vars =>
   638         error ("The environment variable " ^ quote (hd home_vars) ^
   639                " is not set.")
   640       | NONE =>
   641         if File.exists command then
   642           let
   643             (* If slicing is disabled, we expand the last slice to fill the
   644                entire time available. *)
   645             val actual_slices = get_slices slice (best_slices ctxt)
   646             val num_actual_slices = length actual_slices
   647             fun monomorphize_facts facts =
   648               let
   649                 val ctxt =
   650                   ctxt
   651                   |> repair_monomorph_context max_mono_iters best_max_mono_iters
   652                           max_new_mono_instances best_max_new_mono_instances
   653                 (* pseudo-theorem involving the same constants as the subgoal *)
   654                 val subgoal_th =
   655                   Logic.list_implies (hyp_ts, concl_t)
   656                   |> Skip_Proof.make_thm thy
   657                 val rths =
   658                   facts |> chop (length facts div 4)
   659                         |>> map (pair 1 o snd)
   660                         ||> map (pair 2 o snd)
   661                         |> op @
   662                         |> cons (0, subgoal_th)
   663               in
   664                 Monomorph.monomorph atp_schematic_consts_of rths ctxt
   665                 |> fst |> tl
   666                 |> curry ListPair.zip (map fst facts)
   667                 |> maps (fn (name, rths) =>
   668                             map (pair name o zero_var_indexes o snd) rths)
   669               end
   670             fun run_slice time_left (cache_key, cache_value)
   671                     (slice, (time_frac, (complete,
   672                         (key as (best_max_relevant, format, best_type_enc,
   673                                  best_lam_trans, best_uncurried_aliases),
   674                                  extra)))) =
   675               let
   676                 val num_facts =
   677                   length facts |> is_none max_relevant
   678                                   ? Integer.min best_max_relevant
   679                 val soundness = if strict then Strict else Non_Strict
   680                 val type_enc =
   681                   type_enc |> choose_type_enc soundness best_type_enc format
   682                 val sound = is_type_enc_sound type_enc
   683                 val real_ms = Real.fromInt o Time.toMilliseconds
   684                 val slice_timeout =
   685                   ((real_ms time_left
   686                     |> (if slice < num_actual_slices - 1 then
   687                           curry Real.min (time_frac * real_ms timeout)
   688                         else
   689                           I))
   690                    * 0.001) |> seconds
   691                 val generous_slice_timeout =
   692                   Time.+ (slice_timeout, atp_timeout_slack)
   693                 val _ =
   694                   if debug then
   695                     quote name ^ " slice #" ^ string_of_int (slice + 1) ^
   696                     " with " ^ string_of_int num_facts ^ " fact" ^
   697                     plural_s num_facts ^ " for " ^
   698                     string_from_time slice_timeout ^ "..."
   699                     |> Output.urgent_message
   700                   else
   701                     ()
   702                 val readable_names = not (Config.get ctxt atp_full_names)
   703                 val lam_trans =
   704                   case lam_trans of
   705                     SOME s => s
   706                   | NONE => best_lam_trans
   707                 val uncurried_aliases =
   708                   case uncurried_aliases of
   709                     SOME b => b
   710                   | NONE => best_uncurried_aliases
   711                 val value as (atp_problem, _, fact_names, _, _) =
   712                   if cache_key = SOME key then
   713                     cache_value
   714                   else
   715                     facts
   716                     |> map untranslated_fact
   717                     |> not sound
   718                        ? filter_out (is_dangerous_prop ctxt o prop_of o snd)
   719                     |> take num_facts
   720                     |> polymorphism_of_type_enc type_enc <> Raw_Polymorphic
   721                        ? monomorphize_facts
   722                     |> map (apsnd prop_of)
   723                     |> prepare_atp_problem ctxt format prem_role type_enc
   724                                            atp_mode lam_trans uncurried_aliases
   725                                            readable_names true hyp_ts concl_t
   726                 fun sel_weights () = atp_problem_selection_weights atp_problem
   727                 fun ord_info () = atp_problem_term_order_info atp_problem
   728                 val ord = effective_term_order ctxt name
   729                 val full_proof = debug orelse isar_proof
   730                 val args = arguments ctxt full_proof extra slice_timeout
   731                                      (ord, ord_info, sel_weights)
   732                 val command =
   733                   File.shell_path command ^ " " ^ args ^ " " ^
   734                   File.shell_path prob_file
   735                   |> enclose "TIMEFORMAT='%3R'; { time " " ; } 2>&1"
   736                 val _ =
   737                   atp_problem
   738                   |> lines_for_atp_problem format ord ord_info
   739                   |> cons ("% " ^ command ^ "\n")
   740                   |> File.write_list prob_file
   741                 val ((output, run_time), (atp_proof, outcome)) =
   742                   TimeLimit.timeLimit generous_slice_timeout
   743                                       Isabelle_System.bash_output command
   744                   |>> (if overlord then
   745                          prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
   746                        else
   747                          I)
   748                   |> fst |> split_time
   749                   |> (fn accum as (output, _) =>
   750                          (accum,
   751                           extract_tstplike_proof_and_outcome verbose complete
   752                               proof_delims known_failures output
   753                           |>> atp_proof_from_tstplike_proof atp_problem
   754                           handle UNRECOGNIZED_ATP_PROOF () =>
   755                                  ([], SOME ProofIncomplete)))
   756                   handle TimeLimit.TimeOut =>
   757                          (("", slice_timeout), ([], SOME TimedOut))
   758                 val outcome =
   759                   case outcome of
   760                     NONE =>
   761                     (case used_facts_in_unsound_atp_proof ctxt fact_names
   762                                                           atp_proof
   763                           |> Option.map (sort string_ord) of
   764                        SOME facts =>
   765                        let
   766                          val failure =
   767                            UnsoundProof (is_type_enc_sound type_enc, facts)
   768                        in
   769                          if debug then
   770                            (warning (string_for_failure failure); NONE)
   771                          else
   772                            SOME failure
   773                        end
   774                      | NONE => NONE)
   775                   | _ => outcome
   776               in ((SOME key, value), (output, run_time, atp_proof, outcome)) end
   777             val timer = Timer.startRealTimer ()
   778             fun maybe_run_slice slice
   779                     (result as (cache, (_, run_time0, _, SOME _))) =
   780                 let
   781                   val time_left = Time.- (timeout, Timer.checkRealTimer timer)
   782                 in
   783                   if Time.<= (time_left, Time.zeroTime) then
   784                     result
   785                   else
   786                     run_slice time_left cache slice
   787                     |> (fn (cache, (output, run_time, atp_proof, outcome)) =>
   788                            (cache, (output, Time.+ (run_time0, run_time),
   789                                     atp_proof, outcome)))
   790                 end
   791               | maybe_run_slice _ result = result
   792           in
   793             ((NONE, ([], Symtab.empty, Vector.fromList [], [], Symtab.empty)),
   794              ("", Time.zeroTime, [], SOME InternalError))
   795             |> fold maybe_run_slice actual_slices
   796           end
   797         else
   798           error ("Bad executable: " ^ Path.print command)
   799 
   800     (* If the problem file has not been exported, remove it; otherwise, export
   801        the proof file too. *)
   802     fun cleanup prob_file =
   803       if dest_dir = "" then try File.rm prob_file else NONE
   804     fun export prob_file (_, (output, _, _, _)) =
   805       if dest_dir = "" then
   806         ()
   807       else
   808         File.write (Path.explode (Path.implode prob_file ^ "_proof")) output
   809     val ((_, (_, pool, fact_names, _, sym_tab)),
   810          (output, run_time, atp_proof, outcome)) =
   811       with_path cleanup export run_on problem_path_name
   812     val important_message =
   813       if mode = Normal andalso
   814          random_range 0 (atp_important_message_keep_quotient - 1) = 0 then
   815         extract_important_message output
   816       else
   817         ""
   818     val (used_facts, preplay, message, message_tail) =
   819       case outcome of
   820         NONE =>
   821         let
   822           val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof
   823           val needs_full_types = is_typed_helper_used_in_atp_proof atp_proof
   824           val reconstrs =
   825             bunch_of_reconstructors needs_full_types
   826                 (lam_trans_from_atp_proof atp_proof
   827                  o (fn desperate => if desperate then hide_lamsN
   828                                     else metis_default_lam_trans))
   829         in
   830           (used_facts,
   831            fn () =>
   832               let
   833                 val used_pairs =
   834                   facts |> map untranslated_fact |> filter_used_facts used_facts
   835               in
   836                 play_one_line_proof mode debug verbose preplay_timeout
   837                     used_pairs state subgoal (hd reconstrs) reconstrs
   838               end,
   839            fn preplay =>
   840               let
   841                 val isar_params =
   842                   (debug, isar_shrink_factor, pool, fact_names, sym_tab,
   843                    atp_proof, goal)
   844                 val one_line_params =
   845                   (preplay, proof_banner mode name, used_facts,
   846                    choose_minimize_command params minimize_command name preplay,
   847                    subgoal, subgoal_count)
   848               in proof_text ctxt isar_proof isar_params one_line_params end,
   849            (if verbose then
   850               "\nATP real CPU time: " ^ string_from_time run_time ^ "."
   851             else
   852               "") ^
   853            (if important_message <> "" then
   854               "\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^
   855               important_message
   856             else
   857               ""))
   858         end
   859       | SOME failure =>
   860         ([], K (Failed_to_Play plain_metis),
   861          fn _ => string_for_failure failure, "")
   862   in
   863     {outcome = outcome, used_facts = used_facts, run_time = run_time,
   864      preplay = preplay, message = message, message_tail = message_tail}
   865   end
   866 
   867 (* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until
   868    these are sorted out properly in the SMT module, we have to interpret these
   869    ourselves. *)
   870 val remote_smt_failures =
   871   [(2, NoLibwwwPerl),
   872    (22, CantConnect)]
   873 val z3_failures =
   874   [(101, OutOfResources),
   875    (103, MalformedInput),
   876    (110, MalformedInput)]
   877 val unix_failures =
   878   [(139, Crashed)]
   879 val smt_failures = remote_smt_failures @ z3_failures @ unix_failures
   880 
   881 fun failure_from_smt_failure (SMT_Failure.Counterexample {is_real_cex, ...}) =
   882     if is_real_cex then Unprovable else GaveUp
   883   | failure_from_smt_failure SMT_Failure.Time_Out = TimedOut
   884   | failure_from_smt_failure (SMT_Failure.Abnormal_Termination code) =
   885     (case AList.lookup (op =) smt_failures code of
   886        SOME failure => failure
   887      | NONE => UnknownError ("Abnormal termination with exit code " ^
   888                              string_of_int code ^ "."))
   889   | failure_from_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources
   890   | failure_from_smt_failure (SMT_Failure.Other_Failure msg) =
   891     UnknownError msg
   892 
   893 (* FUDGE *)
   894 val smt_max_slices =
   895   Attrib.setup_config_int @{binding sledgehammer_smt_max_slices} (K 8)
   896 val smt_slice_fact_frac =
   897   Attrib.setup_config_real @{binding sledgehammer_smt_slice_fact_frac} (K 0.5)
   898 val smt_slice_time_frac =
   899   Attrib.setup_config_real @{binding sledgehammer_smt_slice_time_frac} (K 0.5)
   900 val smt_slice_min_secs =
   901   Attrib.setup_config_int @{binding sledgehammer_smt_slice_min_secs} (K 5)
   902 
   903 fun smt_filter_loop ctxt name
   904                     ({debug, verbose, overlord, max_mono_iters,
   905                       max_new_mono_instances, timeout, slice, ...} : params)
   906                     state i =
   907   let
   908     val max_slices = if slice then Config.get ctxt smt_max_slices else 1
   909     val repair_context =
   910       select_smt_solver name
   911       #> Config.put SMT_Config.verbose debug
   912       #> (if overlord then
   913             Config.put SMT_Config.debug_files
   914                        (overlord_file_location_for_prover name
   915                         |> (fn (path, name) => path ^ "/" ^ name))
   916           else
   917             I)
   918       #> Config.put SMT_Config.infer_triggers (Config.get ctxt smt_triggers)
   919     val state = state |> Proof.map_context repair_context
   920     fun do_slice timeout slice outcome0 time_so_far facts =
   921       let
   922         val timer = Timer.startRealTimer ()
   923         val state =
   924           state |> Proof.map_context
   925                        (repair_monomorph_context max_mono_iters
   926                             default_max_mono_iters max_new_mono_instances
   927                             default_max_new_mono_instances)
   928         val ms = timeout |> Time.toMilliseconds
   929         val slice_timeout =
   930           if slice < max_slices then
   931             Int.min (ms,
   932                 Int.max (1000 * Config.get ctxt smt_slice_min_secs,
   933                     Real.ceil (Config.get ctxt smt_slice_time_frac
   934                                * Real.fromInt ms)))
   935             |> Time.fromMilliseconds
   936           else
   937             timeout
   938         val num_facts = length facts
   939         val _ =
   940           if debug then
   941             quote name ^ " slice " ^ string_of_int slice ^ " with " ^
   942             string_of_int num_facts ^ " fact" ^ plural_s num_facts ^ " for " ^
   943             string_from_time slice_timeout ^ "..."
   944             |> Output.urgent_message
   945           else
   946             ()
   947         val birth = Timer.checkRealTimer timer
   948         val _ =
   949           if debug then Output.urgent_message "Invoking SMT solver..." else ()
   950         val (outcome, used_facts) =
   951           SMT_Solver.smt_filter_preprocess state facts i
   952           |> SMT_Solver.smt_filter_apply slice_timeout
   953           |> (fn {outcome, used_facts} => (outcome, used_facts))
   954           handle exn => if Exn.is_interrupt exn then
   955                           reraise exn
   956                         else
   957                           (ML_Compiler.exn_message exn
   958                            |> SMT_Failure.Other_Failure |> SOME, [])
   959         val death = Timer.checkRealTimer timer
   960         val outcome0 = if is_none outcome0 then SOME outcome else outcome0
   961         val time_so_far = Time.+ (time_so_far, Time.- (death, birth))
   962         val too_many_facts_perhaps =
   963           case outcome of
   964             NONE => false
   965           | SOME (SMT_Failure.Counterexample _) => false
   966           | SOME SMT_Failure.Time_Out => slice_timeout <> timeout
   967           | SOME (SMT_Failure.Abnormal_Termination _) => true (* kind of *)
   968           | SOME SMT_Failure.Out_Of_Memory => true
   969           | SOME (SMT_Failure.Other_Failure _) => true
   970         val timeout = Time.- (timeout, Timer.checkRealTimer timer)
   971       in
   972         if too_many_facts_perhaps andalso slice < max_slices andalso
   973            num_facts > 0 andalso Time.> (timeout, Time.zeroTime) then
   974           let
   975             val new_num_facts =
   976               Real.ceil (Config.get ctxt smt_slice_fact_frac
   977                          * Real.fromInt num_facts)
   978             val _ =
   979               if verbose andalso is_some outcome then
   980                 quote name ^ " invoked with " ^ string_of_int num_facts ^
   981                 " fact" ^ plural_s num_facts ^ ": " ^
   982                 string_for_failure (failure_from_smt_failure (the outcome)) ^
   983                 " Retrying with " ^ string_of_int new_num_facts ^ " fact" ^
   984                 plural_s new_num_facts ^ "..."
   985                 |> Output.urgent_message
   986               else
   987                 ()
   988           in
   989             facts |> take new_num_facts
   990                   |> do_slice timeout (slice + 1) outcome0 time_so_far
   991           end
   992         else
   993           {outcome = if is_none outcome then NONE else the outcome0,
   994            used_facts = used_facts, run_time = time_so_far}
   995       end
   996   in do_slice timeout 1 NONE Time.zeroTime end
   997 
   998 fun run_smt_solver mode name (params as {debug, verbose, preplay_timeout, ...})
   999         minimize_command
  1000         ({state, subgoal, subgoal_count, facts, ...} : prover_problem) =
  1001   let
  1002     val ctxt = Proof.context_of state
  1003     val num_facts = length facts
  1004     val facts = facts ~~ (0 upto num_facts - 1)
  1005                 |> map (smt_weighted_fact ctxt num_facts)
  1006     val {outcome, used_facts = used_pairs, run_time} =
  1007       smt_filter_loop ctxt name params state subgoal facts
  1008     val used_facts = used_pairs |> map fst
  1009     val outcome = outcome |> Option.map failure_from_smt_failure
  1010     val (preplay, message, message_tail) =
  1011       case outcome of
  1012         NONE =>
  1013         (fn () =>
  1014             play_one_line_proof mode debug verbose preplay_timeout used_pairs
  1015                 state subgoal SMT
  1016                 (bunch_of_reconstructors false
  1017                      (fn plain =>
  1018                          if plain then metis_default_lam_trans else liftingN)),
  1019          fn preplay =>
  1020             let
  1021               val one_line_params =
  1022                 (preplay, proof_banner mode name, used_facts,
  1023                  choose_minimize_command params minimize_command name preplay,
  1024                  subgoal, subgoal_count)
  1025             in one_line_proof_text one_line_params end,
  1026          if verbose then
  1027            "\nSMT solver real CPU time: " ^ string_from_time run_time ^ "."
  1028          else
  1029            "")
  1030       | SOME failure =>
  1031         (K (Failed_to_Play plain_metis), fn _ => string_for_failure failure, "")
  1032   in
  1033     {outcome = outcome, used_facts = used_facts, run_time = run_time,
  1034      preplay = preplay, message = message, message_tail = message_tail}
  1035   end
  1036 
  1037 fun run_reconstructor mode name
  1038         (params as {debug, verbose, timeout, type_enc, lam_trans, ...})
  1039         minimize_command
  1040         ({state, subgoal, subgoal_count, facts, ...} : prover_problem) =
  1041   let
  1042     val reconstr =
  1043       if name = metisN then
  1044         Metis (type_enc |> the_default (hd partial_type_encs),
  1045                lam_trans |> the_default metis_default_lam_trans)
  1046       else if name = smtN then
  1047         SMT
  1048       else
  1049         raise Fail ("unknown reconstructor: " ^ quote name)
  1050     val used_pairs = facts |> map untranslated_fact
  1051     val used_facts = used_pairs |> map fst
  1052   in
  1053     case play_one_line_proof (if mode = Minimize then Normal else mode) debug
  1054                              verbose timeout used_pairs state subgoal reconstr
  1055                              [reconstr] of
  1056       play as Played (_, time) =>
  1057       {outcome = NONE, used_facts = used_facts, run_time = time,
  1058        preplay = K play,
  1059        message =
  1060          fn play =>
  1061             let
  1062               val (_, override_params) = extract_reconstructor params reconstr
  1063               val one_line_params =
  1064                 (play, proof_banner mode name, used_facts,
  1065                  minimize_command override_params name, subgoal,
  1066                  subgoal_count)
  1067             in one_line_proof_text one_line_params end,
  1068        message_tail = ""}
  1069     | play =>
  1070       let
  1071         val failure = case play of Failed_to_Play _ => GaveUp | _ => TimedOut
  1072       in
  1073         {outcome = SOME failure, used_facts = [], run_time = Time.zeroTime,
  1074          preplay = K play, message = fn _ => string_for_failure failure,
  1075          message_tail = ""}
  1076       end
  1077   end
  1078 
  1079 fun get_prover ctxt mode name =
  1080   let val thy = Proof_Context.theory_of ctxt in
  1081     if is_reconstructor name then run_reconstructor mode name
  1082     else if is_atp thy name then run_atp mode name (get_atp thy name ())
  1083     else if is_smt_prover ctxt name then run_smt_solver mode name
  1084     else error ("No such prover: " ^ name ^ ".")
  1085   end
  1086 
  1087 end;