src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML
author blanchet
Wed, 18 Jul 2012 08:44:03 +0200
changeset 49307 7fcee834c7f5
parent 49265 1065c307fafe
child 49308 914ca0827804
permissions -rw-r--r--
more code rationalization in relevance filter
     1 (*  Title:      HOL/Tools/Sledgehammer/sledgehammer_isar.ML
     2     Author:     Jasmin Blanchette, TU Muenchen
     3 
     4 Adds "sledgehammer" and related commands to Isabelle/Isar's outer syntax.
     5 *)
     6 
     7 signature SLEDGEHAMMER_ISAR =
     8 sig
     9   type params = Sledgehammer_Provers.params
    10 
    11   val auto : bool Unsynchronized.ref
    12   val provers : string Unsynchronized.ref
    13   val timeout : int Unsynchronized.ref
    14   val default_params : Proof.context -> (string * string) list -> params
    15   val setup : theory -> theory
    16 end;
    17 
    18 structure Sledgehammer_Isar : SLEDGEHAMMER_ISAR =
    19 struct
    20 
    21 open ATP_Util
    22 open ATP_Systems
    23 open ATP_Problem_Generate
    24 open ATP_Proof_Reconstruct
    25 open Sledgehammer_Util
    26 open Sledgehammer_Fact
    27 open Sledgehammer_Provers
    28 open Sledgehammer_Minimize
    29 open Sledgehammer_Run
    30 
    31 val runN = "run"
    32 val minN = "min"
    33 val messagesN = "messages"
    34 val supported_proversN = "supported_provers"
    35 val running_proversN = "running_provers"
    36 val kill_proversN = "kill_provers"
    37 val refresh_tptpN = "refresh_tptp"
    38 
    39 val auto = Unsynchronized.ref false
    40 
    41 val _ =
    42   ProofGeneralPgip.add_preference Preferences.category_tracing
    43       (Preferences.bool_pref auto "auto-sledgehammer"
    44            "Run Sledgehammer automatically.")
    45 
    46 (** Sledgehammer commands **)
    47 
    48 fun add_fact_override ns : fact_override = {add = ns, del = [], only = false}
    49 fun del_fact_override ns : fact_override = {add = [], del = ns, only = false}
    50 fun only_fact_override ns : fact_override = {add = ns, del = [], only = true}
    51 fun merge_fact_override_pairwise (r1 : fact_override) (r2 : fact_override) =
    52   {add = #add r1 @ #add r2, del = #del r1 @ #del r2,
    53    only = #only r1 andalso #only r2}
    54 fun merge_fact_overrides rs =
    55   fold merge_fact_override_pairwise rs (only_fact_override [])
    56 
    57 (*** parameters ***)
    58 
    59 val provers = Unsynchronized.ref ""
    60 val timeout = Unsynchronized.ref 30
    61 
    62 val _ =
    63   ProofGeneralPgip.add_preference Preferences.category_proof
    64       (Preferences.string_pref provers
    65           "Sledgehammer: Provers"
    66           "Default automatic provers (separated by whitespace)")
    67 
    68 val _ =
    69   ProofGeneralPgip.add_preference Preferences.category_proof
    70       (Preferences.int_pref timeout
    71           "Sledgehammer: Time Limit"
    72           "ATPs will be interrupted after this time (in seconds)")
    73 
    74 type raw_param = string * string list
    75 
    76 val default_default_params =
    77   [("debug", "false"),
    78    ("verbose", "false"),
    79    ("overlord", "false"),
    80    ("blocking", "false"),
    81    ("type_enc", "smart"),
    82    ("strict", "false"),
    83    ("lam_trans", "smart"),
    84    ("uncurried_aliases", "smart"),
    85    ("relevance_thresholds", "0.45 0.85"),
    86    ("max_relevant", "smart"),
    87    ("max_mono_iters", "smart"),
    88    ("max_new_mono_instances", "smart"),
    89    ("isar_proof", "false"),
    90    ("isar_shrink_factor", "1"),
    91    ("slice", "true"),
    92    ("minimize", "smart"),
    93    ("preplay_timeout", "3")]
    94 
    95 val alias_params =
    96   [("prover", ("provers", [])),
    97    ("dont_preplay", ("preplay_timeout", ["0"]))]
    98 val negated_alias_params =
    99   [("no_debug", "debug"),
   100    ("quiet", "verbose"),
   101    ("no_overlord", "overlord"),
   102    ("non_blocking", "blocking"),
   103    ("non_strict", "strict"),
   104    ("no_uncurried_aliases", "uncurried_aliases"),
   105    ("no_isar_proof", "isar_proof"),
   106    ("dont_slice", "slice"),
   107    ("dont_minimize", "minimize")]
   108 
   109 val params_for_minimize =
   110   ["debug", "verbose", "overlord", "type_enc", "strict", "lam_trans",
   111    "uncurried_aliases", "max_mono_iters", "max_new_mono_instances",
   112    "isar_proof", "isar_shrink_factor", "timeout", "preplay_timeout"]
   113 
   114 val property_dependent_params = ["provers", "timeout"]
   115 
   116 fun is_known_raw_param s =
   117   AList.defined (op =) default_default_params s orelse
   118   AList.defined (op =) alias_params s orelse
   119   AList.defined (op =) negated_alias_params s orelse
   120   member (op =) property_dependent_params s orelse s = "expect"
   121 
   122 fun is_prover_list ctxt s =
   123   case space_explode " " s of
   124     ss as _ :: _ => forall (is_prover_supported ctxt) ss
   125   | _ => false
   126 
   127 fun unalias_raw_param (name, value) =
   128   case AList.lookup (op =) alias_params name of
   129     SOME (name', []) => (name', value)
   130   | SOME (param' as (name', _)) =>
   131     if value <> ["false"] then
   132       param'
   133     else
   134       error ("Parameter " ^ quote name ^ " cannot be set to \"false\" \
   135              \(cf. " ^ quote name' ^ ").")
   136   | NONE =>
   137     case AList.lookup (op =) negated_alias_params name of
   138       SOME name' => (name', case value of
   139                               ["false"] => ["true"]
   140                             | ["true"] => ["false"]
   141                             | [] => ["false"]
   142                             | _ => value)
   143     | NONE => (name, value)
   144 
   145 val any_type_enc = type_enc_from_string Strict "erased"
   146 
   147 (* "provers =", "type_enc =", and "lam_trans" can be omitted. For the last two,
   148    this is a secret feature. *)
   149 fun normalize_raw_param ctxt =
   150   unalias_raw_param
   151   #> (fn (name, value) =>
   152          if is_known_raw_param name then
   153            (name, value)
   154          else if is_prover_list ctxt name andalso null value then
   155            ("provers", [name])
   156          else if can (type_enc_from_string Strict) name andalso null value then
   157            ("type_enc", [name])
   158          else if can (trans_lams_from_string ctxt any_type_enc) name andalso
   159                  null value then
   160            ("lam_trans", [name])
   161          else
   162            error ("Unknown parameter: " ^ quote name ^ "."))
   163 
   164 
   165 (* Ensures that type encodings such as "mono_native?" and "poly_guards!!" are
   166    read correctly. *)
   167 val implode_param = strip_spaces_except_between_idents o space_implode " "
   168 
   169 structure Data = Theory_Data
   170 (
   171   type T = raw_param list
   172   val empty = default_default_params |> map (apsnd single)
   173   val extend = I
   174   fun merge data : T = AList.merge (op =) (K true) data
   175 )
   176 
   177 fun remotify_prover_if_supported_and_not_already_remote ctxt name =
   178   if String.isPrefix remote_prefix name then
   179     SOME name
   180   else
   181     let val remote_name = remote_prefix ^ name in
   182       if is_prover_supported ctxt remote_name then SOME remote_name else NONE
   183     end
   184 
   185 fun remotify_prover_if_not_installed ctxt name =
   186   if is_prover_supported ctxt name andalso is_prover_installed ctxt name then
   187     SOME name
   188   else
   189     remotify_prover_if_supported_and_not_already_remote ctxt name
   190 
   191 fun avoid_too_many_threads _ _ [] = []
   192   | avoid_too_many_threads _ (0, 0) _ = []
   193   | avoid_too_many_threads ctxt (0, max_remote) provers =
   194     provers
   195     |> map_filter (remotify_prover_if_supported_and_not_already_remote ctxt)
   196     |> take max_remote
   197   | avoid_too_many_threads _ (max_local, 0) provers =
   198     provers
   199     |> filter_out (String.isPrefix remote_prefix)
   200     |> take max_local
   201   | avoid_too_many_threads ctxt max_local_and_remote (prover :: provers) =
   202     let
   203       val max_local_and_remote =
   204         max_local_and_remote
   205         |> (if String.isPrefix remote_prefix prover then apsnd else apfst)
   206               (Integer.add ~1)
   207     in prover :: avoid_too_many_threads ctxt max_local_and_remote provers end
   208 
   209 val max_default_remote_threads = 4
   210 
   211 (* The first ATP of the list is used by Auto Sledgehammer. Because of the low
   212    timeout, it makes sense to put SPASS first. *)
   213 fun default_provers_param_value ctxt =
   214   [spassN, eN, vampireN, SMT_Solver.solver_name_of ctxt, e_sineN,
   215    waldmeisterN]
   216   |> map_filter (remotify_prover_if_not_installed ctxt)
   217   |> avoid_too_many_threads ctxt (Multithreading.max_threads_value (),
   218                                   max_default_remote_threads)
   219   |> implode_param
   220 
   221 fun set_default_raw_param param thy =
   222   let val ctxt = Proof_Context.init_global thy in
   223     thy |> Data.map (AList.update (op =) (normalize_raw_param ctxt param))
   224   end
   225 fun default_raw_params ctxt =
   226   let val thy = Proof_Context.theory_of ctxt in
   227     Data.get thy
   228     |> fold (AList.default (op =))
   229             [("provers", [case !provers of
   230                             "" => default_provers_param_value ctxt
   231                           | s => s]),
   232              ("timeout", let val timeout = !timeout in
   233                            [if timeout <= 0 then "none"
   234                             else string_of_int timeout]
   235                          end)]
   236   end
   237 
   238 val infinity_time_in_secs = 60 * 60 * 24 * 365
   239 val the_timeout = the_default (Time.fromSeconds infinity_time_in_secs)
   240 
   241 fun extract_params mode default_params override_params =
   242   let
   243     val raw_params = rev override_params @ rev default_params
   244     val lookup = Option.map implode_param o AList.lookup (op =) raw_params
   245     val lookup_string = the_default "" o lookup
   246     fun general_lookup_bool option default_value name =
   247       case lookup name of
   248         SOME s => parse_bool_option option name s
   249       | NONE => default_value
   250     val lookup_bool = the o general_lookup_bool false (SOME false)
   251     fun lookup_time name =
   252       case lookup name of
   253         SOME s => parse_time_option name s
   254       | NONE => NONE
   255     fun lookup_int name =
   256       case lookup name of
   257         NONE => 0
   258       | SOME s => case Int.fromString s of
   259                     SOME n => n
   260                   | NONE => error ("Parameter " ^ quote name ^
   261                                    " must be assigned an integer value.")
   262     fun lookup_real_pair name =
   263       case lookup name of
   264         NONE => (0.0, 0.0)
   265       | SOME s => case s |> space_explode " " |> map Real.fromString of
   266                     [SOME r1, SOME r2] => (r1, r2)
   267                   | _ => error ("Parameter " ^ quote name ^
   268                                 "must be assigned a pair of floating-point \
   269                                 \values (e.g., \"0.6 0.95\")")
   270     fun lookup_option lookup' name =
   271       case lookup name of
   272         SOME "smart" => NONE
   273       | _ => SOME (lookup' name)
   274     val debug = mode <> Auto_Try andalso lookup_bool "debug"
   275     val verbose = debug orelse (mode <> Auto_Try andalso lookup_bool "verbose")
   276     val overlord = lookup_bool "overlord"
   277     val blocking =
   278       Isabelle_Process.is_active () orelse mode <> Normal orelse debug orelse
   279       lookup_bool "blocking"
   280     val provers = lookup_string "provers" |> space_explode " "
   281                   |> mode = Auto_Try ? single o hd
   282     val type_enc =
   283       if mode = Auto_Try then
   284         NONE
   285       else case lookup_string "type_enc" of
   286         "smart" => NONE
   287       | s => (type_enc_from_string Strict s; SOME s)
   288     val strict = mode = Auto_Try orelse lookup_bool "strict"
   289     val lam_trans = lookup_option lookup_string "lam_trans"
   290     val uncurried_aliases = lookup_option lookup_bool "uncurried_aliases"
   291     val relevance_thresholds = lookup_real_pair "relevance_thresholds"
   292     val max_relevant = lookup_option lookup_int "max_relevant"
   293     val max_mono_iters = lookup_option lookup_int "max_mono_iters"
   294     val max_new_mono_instances =
   295       lookup_option lookup_int "max_new_mono_instances"
   296     val isar_proof = lookup_bool "isar_proof"
   297     val isar_shrink_factor = Int.max (1, lookup_int "isar_shrink_factor")
   298     val slice = mode <> Auto_Try andalso lookup_bool "slice"
   299     val minimize =
   300       if mode = Auto_Try then NONE else lookup_option lookup_bool "minimize"
   301     val timeout =
   302       (if mode = Auto_Try then NONE else lookup_time "timeout") |> the_timeout
   303     val preplay_timeout =
   304       if mode = Auto_Try then Time.zeroTime
   305       else lookup_time "preplay_timeout" |> the_timeout
   306     val expect = lookup_string "expect"
   307   in
   308     {debug = debug, verbose = verbose, overlord = overlord, blocking = blocking,
   309      provers = provers, type_enc = type_enc, strict = strict,
   310      lam_trans = lam_trans, uncurried_aliases = uncurried_aliases,
   311      relevance_thresholds = relevance_thresholds, max_relevant = max_relevant,
   312      max_mono_iters = max_mono_iters,
   313      max_new_mono_instances = max_new_mono_instances,  isar_proof = isar_proof,
   314      isar_shrink_factor = isar_shrink_factor, slice = slice,
   315      minimize = minimize, timeout = timeout, preplay_timeout = preplay_timeout,
   316      expect = expect}
   317   end
   318 
   319 fun get_params mode = extract_params mode o default_raw_params
   320 fun default_params thy = get_params Normal thy o map (apsnd single)
   321 
   322 (* Sledgehammer the given subgoal *)
   323 
   324 val default_minimize_prover = metisN
   325 
   326 fun is_raw_param_relevant_for_minimize (name, _) =
   327   member (op =) params_for_minimize name
   328 fun string_for_raw_param (key, values) =
   329   key ^ (case implode_param values of "" => "" | value => " = " ^ value)
   330 
   331 fun minimize_command override_params i more_override_params prover_name
   332                      fact_names =
   333   let
   334     val params =
   335       (override_params
   336        |> filter_out (AList.defined (op =) more_override_params o fst)) @
   337       more_override_params
   338       |> filter is_raw_param_relevant_for_minimize
   339       |> map string_for_raw_param
   340       |> (if prover_name = default_minimize_prover then I else cons prover_name)
   341       |> space_implode ", "
   342   in
   343     "sledgehammer" ^ " " ^ minN ^
   344     (if params = "" then "" else enclose " [" "]" params) ^
   345     " (" ^ space_implode " " fact_names ^ ")" ^
   346     (if i = 1 then "" else " " ^ string_of_int i)
   347   end
   348 
   349 fun hammer_away override_params subcommand opt_i fact_override state =
   350   let
   351     val ctxt = Proof.context_of state
   352     val override_params = override_params |> map (normalize_raw_param ctxt)
   353   in
   354     if subcommand = runN then
   355       let val i = the_default 1 opt_i in
   356         run_sledgehammer (get_params Normal ctxt override_params) Normal i
   357                          fact_override (minimize_command override_params i)
   358                          state
   359         |> K ()
   360       end
   361     else if subcommand = minN then
   362       run_minimize (get_params Minimize ctxt
   363                                (("provers", [default_minimize_prover]) ::
   364                                 override_params))
   365                    (the_default 1 opt_i) (#add fact_override) state
   366     else if subcommand = messagesN then
   367       messages opt_i
   368     else if subcommand = supported_proversN then
   369       supported_provers ctxt
   370     else if subcommand = running_proversN then
   371       running_provers ()
   372     else if subcommand = kill_proversN then
   373       kill_provers ()
   374     else if subcommand = refresh_tptpN then
   375       refresh_systems_on_tptp ()
   376     else
   377       error ("Unknown subcommand: " ^ quote subcommand ^ ".")
   378   end
   379 
   380 fun sledgehammer_trans (((subcommand, params), fact_override), opt_i) =
   381   Toplevel.keep (hammer_away params subcommand opt_i fact_override
   382                  o Toplevel.proof_of)
   383 
   384 fun string_for_raw_param (name, value) = name ^ " = " ^ implode_param value
   385 
   386 fun sledgehammer_params_trans params =
   387   Toplevel.theory
   388       (fold set_default_raw_param params
   389        #> tap (fn thy =>
   390                   let val ctxt = Proof_Context.init_global thy in
   391                     writeln ("Default parameters for Sledgehammer:\n" ^
   392                              (case default_raw_params ctxt |> rev of
   393                                 [] => "none"
   394                               | params =>
   395                                 params |> map string_for_raw_param
   396                                        |> sort_strings |> cat_lines))
   397                   end))
   398 
   399 val parse_query_bang = @{keyword "?"} || @{keyword "!"} || @{keyword "!!"}
   400 val parse_key =
   401   Scan.repeat1 (Parse.typ_group || parse_query_bang) >> implode_param
   402 val parse_value =
   403   Scan.repeat1 (Parse.xname || Parse.float_number || parse_query_bang)
   404 val parse_param = parse_key -- Scan.optional (@{keyword "="} |-- parse_value) []
   405 val parse_params = Scan.optional (Args.bracks (Parse.list parse_param)) []
   406 val parse_fact_refs =
   407   Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) Parse_Spec.xthm)
   408 val parse_relevance_chunk =
   409   (Args.add |-- Args.colon |-- parse_fact_refs >> add_fact_override)
   410   || (Args.del |-- Args.colon |-- parse_fact_refs >> del_fact_override)
   411   || (parse_fact_refs >> only_fact_override)
   412 val parse_fact_override =
   413   Scan.optional (Args.parens (Scan.repeat parse_relevance_chunk
   414                               >> merge_fact_overrides))
   415                 no_fact_override
   416 
   417 val _ =
   418   Outer_Syntax.improper_command @{command_spec "sledgehammer"}
   419     "search for first-order proof using automatic theorem provers"
   420     ((Scan.optional Parse.short_ident runN -- parse_params
   421       -- parse_fact_override -- Scan.option Parse.nat) #>> sledgehammer_trans)
   422 val _ =
   423   Outer_Syntax.command @{command_spec "sledgehammer_params"}
   424     "set and display the default parameters for Sledgehammer"
   425     (parse_params #>> sledgehammer_params_trans)
   426 
   427 fun try_sledgehammer auto state =
   428   let
   429     val ctxt = Proof.context_of state
   430     val mode = if auto then Auto_Try else Try
   431     val i = 1
   432   in
   433     run_sledgehammer (get_params mode ctxt []) mode i no_fact_override
   434                      (minimize_command [] i) state
   435   end
   436 
   437 val setup = Try.register_tool ("sledgehammer", (40, auto, try_sledgehammer))
   438 
   439 end;