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