src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML
author blanchet
Wed, 03 Nov 2010 22:26:53 +0100
changeset 40582 03156257040f
parent 40422 3788b7adab36
child 40584 4521d56aef63
permissions -rw-r--r--
standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet@36375
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_isar.ML
blanchet@35866
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@35866
     3
blanchet@35866
     4
Adds "sledgehammer" and related commands to Isabelle/Isar's outer syntax.
blanchet@35866
     5
*)
blanchet@35866
     6
blanchet@35961
     7
signature SLEDGEHAMMER_ISAR =
blanchet@35961
     8
sig
blanchet@38255
     9
  type params = Sledgehammer.params
blanchet@35961
    10
blanchet@39564
    11
  val auto : bool Unsynchronized.ref
blanchet@40240
    12
  val provers : string Unsynchronized.ref
blanchet@39564
    13
  val timeout : int Unsynchronized.ref
blanchet@39564
    14
  val full_types : bool Unsynchronized.ref
blanchet@40250
    15
  val default_params : Proof.context -> (string * string) list -> params
blanchet@39564
    16
  val setup : theory -> theory
blanchet@35961
    17
end;
blanchet@35961
    18
blanchet@35969
    19
structure Sledgehammer_Isar : SLEDGEHAMMER_ISAR =
blanchet@35866
    20
struct
blanchet@35866
    21
blanchet@38262
    22
open ATP_Systems
blanchet@35969
    23
open Sledgehammer_Util
blanchet@38255
    24
open Sledgehammer
blanchet@39232
    25
open Sledgehammer_Minimize
blanchet@35866
    26
blanchet@39564
    27
val auto = Unsynchronized.ref false
blanchet@39564
    28
blanchet@39564
    29
val _ =
blanchet@39564
    30
  ProofGeneralPgip.add_preference Preferences.category_tracing
blanchet@39564
    31
      (Preferences.bool_pref auto "auto-sledgehammer"
blanchet@39573
    32
           "Run Sledgehammer automatically.")
blanchet@39564
    33
blanchet@36394
    34
(** Sledgehammer commands **)
blanchet@36394
    35
blanchet@39564
    36
val no_relevance_override = {add = [], del = [], only = false}
blanchet@39564
    37
fun add_relevance_override ns : relevance_override =
blanchet@35964
    38
  {add = ns, del = [], only = false}
blanchet@39564
    39
fun del_relevance_override ns : relevance_override =
blanchet@35964
    40
  {add = [], del = ns, only = false}
blanchet@35964
    41
fun only_relevance_override ns : relevance_override =
blanchet@35964
    42
  {add = ns, del = [], only = true}
blanchet@36003
    43
fun merge_relevance_override_pairwise (r1 : relevance_override)
blanchet@36003
    44
                                      (r2 : relevance_override) =
blanchet@35964
    45
  {add = #add r1 @ #add r2, del = #del r1 @ #del r2,
blanchet@36183
    46
   only = #only r1 andalso #only r2}
blanchet@36188
    47
fun merge_relevance_overrides rs =
blanchet@36188
    48
  fold merge_relevance_override_pairwise rs (only_relevance_override [])
blanchet@35963
    49
blanchet@36371
    50
(*** parameters ***)
blanchet@36371
    51
blanchet@40240
    52
val provers = Unsynchronized.ref ""
blanchet@39581
    53
val timeout = Unsynchronized.ref 30
blanchet@36371
    54
val full_types = Unsynchronized.ref false
blanchet@36371
    55
blanchet@36371
    56
val _ =
blanchet@36371
    57
  ProofGeneralPgip.add_preference Preferences.category_proof
blanchet@40240
    58
      (Preferences.string_pref provers
blanchet@40240
    59
          "Sledgehammer: Provers"
blanchet@36371
    60
          "Default automatic provers (separated by whitespace)")
blanchet@36371
    61
blanchet@36371
    62
val _ =
blanchet@36371
    63
  ProofGeneralPgip.add_preference Preferences.category_proof
blanchet@36371
    64
      (Preferences.int_pref timeout
blanchet@36373
    65
          "Sledgehammer: Time Limit"
blanchet@36371
    66
          "ATPs will be interrupted after this time (in seconds)")
blanchet@36371
    67
blanchet@36371
    68
val _ =
blanchet@36371
    69
  ProofGeneralPgip.add_preference Preferences.category_proof
blanchet@36371
    70
      (Preferences.bool_pref full_types
blanchet@36373
    71
          "Sledgehammer: Full Types" "ATPs will use full type information")
blanchet@36371
    72
blanchet@35961
    73
type raw_param = string * string list
blanchet@35866
    74
blanchet@35961
    75
val default_default_params =
blanchet@39226
    76
  [("blocking", "false"),
blanchet@39226
    77
   ("debug", "false"),
blanchet@35961
    78
   ("verbose", "false"),
blanchet@36141
    79
   ("overlord", "false"),
blanchet@36235
    80
   ("explicit_apply", "false"),
blanchet@39165
    81
   ("relevance_thresholds", "45 85"),
blanchet@38983
    82
   ("max_relevant", "smart"),
blanchet@35961
    83
   ("isar_proof", "false"),
blanchet@38813
    84
   ("isar_shrink_factor", "1")]
blanchet@35866
    85
blanchet@36138
    86
val alias_params =
blanchet@40240
    87
  [("prover", "provers"),
blanchet@40240
    88
   ("atps", "provers"), (* FIXME: legacy *)
blanchet@40240
    89
   ("atp", "provers")]  (* FIXME: legacy *)
blanchet@36138
    90
val negated_alias_params =
blanchet@39226
    91
  [("non_blocking", "blocking"),
blanchet@39226
    92
   ("no_debug", "debug"),
blanchet@35961
    93
   ("quiet", "verbose"),
blanchet@36141
    94
   ("no_overlord", "overlord"),
blanchet@36399
    95
   ("partial_types", "full_types"),
blanchet@36235
    96
   ("implicit_apply", "explicit_apply"),
blanchet@36900
    97
   ("no_isar_proof", "isar_proof")]
blanchet@35866
    98
blanchet@36281
    99
val params_for_minimize =
blanchet@38346
   100
  ["debug", "verbose", "overlord", "full_types", "isar_proof",
blanchet@38813
   101
   "isar_shrink_factor", "timeout"]
blanchet@36281
   102
blanchet@40240
   103
val property_dependent_params = ["provers", "full_types", "timeout"]
blanchet@35866
   104
blanchet@35961
   105
fun is_known_raw_param s =
blanchet@35961
   106
  AList.defined (op =) default_default_params s orelse
blanchet@36138
   107
  AList.defined (op =) alias_params s orelse
blanchet@36138
   108
  AList.defined (op =) negated_alias_params s orelse
blanchet@39229
   109
  member (op =) property_dependent_params s orelse s = "expect"
blanchet@35866
   110
blanchet@35961
   111
fun check_raw_param (s, _) =
blanchet@35961
   112
  if is_known_raw_param s then ()
blanchet@35961
   113
  else error ("Unknown parameter: " ^ quote s ^ ".")
blanchet@35961
   114
blanchet@36138
   115
fun unalias_raw_param (name, value) =
blanchet@36138
   116
  case AList.lookup (op =) alias_params name of
blanchet@36138
   117
    SOME name' => (name', value)
blanchet@36138
   118
  | NONE =>
blanchet@36138
   119
    case AList.lookup (op =) negated_alias_params name of
blanchet@36138
   120
      SOME name' => (name', case value of
blanchet@36138
   121
                              ["false"] => ["true"]
blanchet@36138
   122
                            | ["true"] => ["false"]
blanchet@36138
   123
                            | [] => ["false"]
blanchet@36138
   124
                            | _ => value)
blanchet@36138
   125
    | NONE => (name, value)
blanchet@35961
   126
blanchet@35961
   127
structure Data = Theory_Data(
blanchet@35961
   128
  type T = raw_param list
blanchet@35961
   129
  val empty = default_default_params |> map (apsnd single)
blanchet@35961
   130
  val extend = I
blanchet@35961
   131
  fun merge p : T = AList.merge (op =) (K true) p)
blanchet@35961
   132
blanchet@40253
   133
fun remotify_prover_if_available_and_not_already_remote thy name =
blanchet@40253
   134
  if String.isPrefix remote_prefix name then
blanchet@40253
   135
    SOME name
blanchet@40253
   136
  else
blanchet@40253
   137
    let val remote_name = remote_prefix ^ name in
blanchet@40253
   138
      if is_prover_available thy remote_name then SOME remote_name else NONE
blanchet@40253
   139
    end
blanchet@40253
   140
blanchet@40253
   141
fun remotify_prover_if_not_installed ctxt name =
blanchet@40253
   142
  let val thy = ProofContext.theory_of ctxt in
blanchet@40253
   143
    if is_prover_available thy name andalso is_prover_installed ctxt name then
blanchet@40253
   144
      SOME name
blanchet@40253
   145
    else
blanchet@40253
   146
      remotify_prover_if_available_and_not_already_remote thy name
blanchet@40253
   147
  end
blanchet@40253
   148
blanchet@40253
   149
fun avoid_too_many_local_threads _ _ [] = []
blanchet@40253
   150
  | avoid_too_many_local_threads thy 0 provers =
blanchet@40253
   151
    map_filter (remotify_prover_if_available_and_not_already_remote thy) provers
blanchet@40253
   152
  | avoid_too_many_local_threads thy n (prover :: provers) =
blanchet@40253
   153
    let val n = if String.isPrefix remote_prefix prover then n else n - 1 in
blanchet@40253
   154
      prover :: avoid_too_many_local_threads thy n provers
blanchet@40253
   155
    end
blanchet@40240
   156
blanchet@40385
   157
fun num_processors () = try Thread.numProcessors () |> the_default 1
blanchet@40359
   158
blanchet@40240
   159
(* The first ATP of the list is used by Auto Sledgehammer. Because of the low
blanchet@40240
   160
   timeout, it makes sense to put SPASS first. *)
blanchet@40250
   161
fun default_provers_param_value ctxt =
blanchet@40250
   162
  let val thy = ProofContext.theory_of ctxt in
blanchet@40253
   163
    [spassN, eN, vampireN, sine_eN (* FIXME: , smtN *)]
blanchet@40253
   164
    |> map_filter (remotify_prover_if_not_installed ctxt)
blanchet@40359
   165
    |> avoid_too_many_local_threads thy (num_processors ())
blanchet@40250
   166
    |> space_implode " "
blanchet@40250
   167
  end
blanchet@40240
   168
blanchet@36138
   169
val set_default_raw_param = Data.map o AList.update (op =) o unalias_raw_param
blanchet@40250
   170
fun default_raw_params ctxt =
blanchet@40250
   171
  let val thy = ProofContext.theory_of ctxt in
blanchet@40250
   172
    Data.get thy
blanchet@40250
   173
    |> fold (AList.default (op =))
blanchet@40250
   174
            [("provers", [case !provers of
blanchet@40250
   175
                            "" => default_provers_param_value ctxt
blanchet@40250
   176
                          | s => s]),
blanchet@40250
   177
             ("full_types", [if !full_types then "true" else "false"]),
blanchet@40250
   178
             ("timeout", let val timeout = !timeout in
blanchet@40250
   179
                           [if timeout <= 0 then "none"
blanchet@40582
   180
                            else string_of_int timeout]
blanchet@40250
   181
                         end)]
blanchet@40250
   182
  end
blanchet@35961
   183
blanchet@35961
   184
val infinity_time_in_secs = 60 * 60 * 24 * 365
blanchet@35961
   185
val the_timeout = the_default (Time.fromSeconds infinity_time_in_secs)
blanchet@35961
   186
blanchet@39564
   187
fun extract_params auto default_params override_params =
blanchet@35961
   188
  let
blanchet@36138
   189
    val override_params = map unalias_raw_param override_params
blanchet@35961
   190
    val raw_params = rev override_params @ rev default_params
blanchet@35961
   191
    val lookup = Option.map (space_implode " ") o AList.lookup (op =) raw_params
blanchet@35961
   192
    val lookup_string = the_default "" o lookup
blanchet@35961
   193
    fun general_lookup_bool option default_value name =
blanchet@35961
   194
      case lookup name of
blanchet@35968
   195
        SOME s => parse_bool_option option name s
blanchet@35961
   196
      | NONE => default_value
blanchet@35961
   197
    val lookup_bool = the o general_lookup_bool false (SOME false)
blanchet@35961
   198
    fun lookup_time name =
blanchet@39564
   199
      case lookup name of
blanchet@39564
   200
        SOME s => parse_time_option name s
blanchet@39564
   201
      | NONE => NONE
blanchet@35964
   202
    fun lookup_int name =
blanchet@35961
   203
      case lookup name of
blanchet@35964
   204
        NONE => 0
blanchet@35964
   205
      | SOME s => case Int.fromString s of
blanchet@35964
   206
                    SOME n => n
blanchet@35964
   207
                  | NONE => error ("Parameter " ^ quote name ^
blanchet@35964
   208
                                   " must be assigned an integer value.")
blanchet@38984
   209
    fun lookup_int_pair name =
blanchet@38984
   210
      case lookup name of
blanchet@38984
   211
        NONE => (0, 0)
blanchet@38984
   212
      | SOME s => case s |> space_explode " " |> map Int.fromString of
blanchet@38984
   213
                    [SOME n1, SOME n2] => (n1, n2)
blanchet@38984
   214
                  | _ => error ("Parameter " ^ quote name ^
blanchet@38984
   215
                                "must be assigned a pair of integer values \
blanchet@38984
   216
                                \(e.g., \"60 95\")")
blanchet@38812
   217
    fun lookup_int_option name =
blanchet@38812
   218
      case lookup name of
blanchet@38812
   219
        SOME "smart" => NONE
blanchet@38812
   220
      | _ => SOME (lookup_int name)
blanchet@39564
   221
    val blocking = auto orelse lookup_bool "blocking"
blanchet@39564
   222
    val debug = not auto andalso lookup_bool "debug"
blanchet@39564
   223
    val verbose = debug orelse (not auto andalso lookup_bool "verbose")
blanchet@36141
   224
    val overlord = lookup_bool "overlord"
blanchet@40240
   225
    val provers = lookup_string "provers" |> space_explode " "
blanchet@40240
   226
                  |> auto ? single o hd
blanchet@35961
   227
    val full_types = lookup_bool "full_types"
blanchet@36235
   228
    val explicit_apply = lookup_bool "explicit_apply"
blanchet@38984
   229
    val relevance_thresholds =
blanchet@38984
   230
      lookup_int_pair "relevance_thresholds"
blanchet@38984
   231
      |> pairself (fn n => 0.01 * Real.fromInt n)
blanchet@38983
   232
    val max_relevant = lookup_int_option "max_relevant"
blanchet@35961
   233
    val isar_proof = lookup_bool "isar_proof"
blanchet@36916
   234
    val isar_shrink_factor = Int.max (1, lookup_int "isar_shrink_factor")
blanchet@39564
   235
    val timeout = (if auto then NONE else lookup_time "timeout") |> the_timeout
blanchet@39229
   236
    val expect = lookup_string "expect"
blanchet@35961
   237
  in
blanchet@39226
   238
    {blocking = blocking, debug = debug, verbose = verbose, overlord = overlord,
blanchet@40240
   239
     provers = provers, full_types = full_types,
blanchet@40240
   240
     explicit_apply = explicit_apply,
blanchet@38984
   241
     relevance_thresholds = relevance_thresholds, max_relevant = max_relevant,
blanchet@39241
   242
     isar_proof = isar_proof, isar_shrink_factor = isar_shrink_factor,
blanchet@39241
   243
     timeout = timeout, expect = expect}
blanchet@35961
   244
  end
blanchet@35961
   245
blanchet@40250
   246
fun get_params auto ctxt = extract_params auto (default_raw_params ctxt)
blanchet@39564
   247
fun default_params thy = get_params false thy o map (apsnd single)
blanchet@35866
   248
blanchet@36373
   249
(* Sledgehammer the given subgoal *)
blanchet@36373
   250
blanchet@36373
   251
val sledgehammerN = "sledgehammer"
blanchet@36373
   252
val sledgehammer_paramsN = "sledgehammer_params"
blanchet@36373
   253
blanchet@35963
   254
val runN = "run"
blanchet@35963
   255
val minimizeN = "minimize"
blanchet@35963
   256
val messagesN = "messages"
blanchet@40240
   257
val available_proversN = "available_provers"
blanchet@40240
   258
val running_proversN = "running_provers"
blanchet@40240
   259
val kill_proversN = "kill_provers"
blanchet@35963
   260
val refresh_tptpN = "refresh_tptp"
blanchet@35866
   261
blanchet@36281
   262
val is_raw_param_relevant_for_minimize =
blanchet@36281
   263
  member (op =) params_for_minimize o fst o unalias_raw_param
blanchet@36281
   264
fun string_for_raw_param (key, values) =
blanchet@36477
   265
  key ^ (case space_implode " " values of "" => "" | value => " = " ^ value)
blanchet@36281
   266
blanchet@40240
   267
fun minimize_command override_params i prover_name fact_names =
blanchet@40240
   268
  sledgehammerN ^ " " ^ minimizeN ^ " [prover = " ^ prover_name ^
blanchet@36281
   269
  (override_params |> filter is_raw_param_relevant_for_minimize
blanchet@36281
   270
                   |> implode o map (prefix ", " o string_for_raw_param)) ^
blanchet@37165
   271
  "] (" ^ space_implode " " fact_names ^ ")" ^
blanchet@36281
   272
  (if i = 1 then "" else " " ^ string_of_int i)
blanchet@36281
   273
blanchet@35964
   274
fun hammer_away override_params subcommand opt_i relevance_override state =
blanchet@35961
   275
  let
blanchet@40250
   276
    val ctxt = Proof.context_of state
blanchet@35961
   277
    val thy = Proof.theory_of state
blanchet@36477
   278
    val _ = app check_raw_param override_params
blanchet@35963
   279
  in
blanchet@35963
   280
    if subcommand = runN then
blanchet@36281
   281
      let val i = the_default 1 opt_i in
blanchet@40250
   282
        run_sledgehammer (get_params false ctxt override_params) false i
blanchet@39564
   283
                         relevance_override (minimize_command override_params i)
blanchet@39564
   284
                         state
blanchet@39564
   285
        |> K ()
blanchet@36281
   286
      end
blanchet@35963
   287
    else if subcommand = minimizeN then
blanchet@40250
   288
      run_minimize (get_params false ctxt override_params) (the_default 1 opt_i)
blanchet@38813
   289
                   (#add relevance_override) state
blanchet@35963
   290
    else if subcommand = messagesN then
blanchet@35963
   291
      messages opt_i
blanchet@40240
   292
    else if subcommand = available_proversN then
blanchet@40240
   293
      available_provers thy
blanchet@40240
   294
    else if subcommand = running_proversN then
blanchet@40240
   295
      running_provers ()
blanchet@40240
   296
    else if subcommand = kill_proversN then
blanchet@40240
   297
      kill_provers ()
blanchet@35963
   298
    else if subcommand = refresh_tptpN then
blanchet@35963
   299
      refresh_systems_on_tptp ()
blanchet@35963
   300
    else
blanchet@35963
   301
      error ("Unknown subcommand: " ^ quote subcommand ^ ".")
blanchet@35963
   302
  end
blanchet@35961
   303
blanchet@35964
   304
fun sledgehammer_trans (((subcommand, params), relevance_override), opt_i) =
blanchet@35964
   305
  Toplevel.keep (hammer_away params subcommand opt_i relevance_override
blanchet@35964
   306
                 o Toplevel.proof_of)
blanchet@35961
   307
blanchet@35961
   308
fun string_for_raw_param (name, value) = name ^ " = " ^ space_implode " " value
blanchet@35961
   309
blanchet@35963
   310
fun sledgehammer_params_trans params =
blanchet@35961
   311
  Toplevel.theory
blanchet@35963
   312
      (fold set_default_raw_param params
blanchet@39564
   313
       #> tap (fn thy =>
blanchet@40250
   314
                  let val ctxt = ProofContext.init_global thy in
blanchet@40250
   315
                    writeln ("Default parameters for Sledgehammer:\n" ^
blanchet@40250
   316
                             (case default_raw_params ctxt |> rev of
blanchet@40250
   317
                                [] => "none"
blanchet@40250
   318
                              | params =>
blanchet@40250
   319
                                (map check_raw_param params;
blanchet@40250
   320
                                 params |> map string_for_raw_param
blanchet@40250
   321
                                        |> sort_strings |> cat_lines)))
blanchet@40250
   322
                  end))
blanchet@35961
   323
wenzelm@36970
   324
val parse_key = Scan.repeat1 Parse.typ_group >> space_implode " "
blanchet@40582
   325
val parse_value = Scan.repeat1 (Parse.xname || Parse.float_number)
wenzelm@36970
   326
val parse_param = parse_key -- Scan.optional (Parse.$$$ "=" |-- parse_value) []
wenzelm@36970
   327
val parse_params = Scan.optional (Args.bracks (Parse.list parse_param)) []
blanchet@35963
   328
val parse_fact_refs =
blanchet@39240
   329
  Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) Parse_Spec.xthm)
blanchet@35964
   330
val parse_relevance_chunk =
blanchet@39564
   331
  (Args.add |-- Args.colon |-- parse_fact_refs >> add_relevance_override)
blanchet@39564
   332
  || (Args.del |-- Args.colon |-- parse_fact_refs >> del_relevance_override)
blanchet@36188
   333
  || (parse_fact_refs >> only_relevance_override)
blanchet@35964
   334
val parse_relevance_override =
blanchet@36188
   335
  Scan.optional (Args.parens (Scan.repeat parse_relevance_chunk
blanchet@36188
   336
                              >> merge_relevance_overrides))
blanchet@39564
   337
                no_relevance_override
blanchet@35961
   338
val parse_sledgehammer_command =
blanchet@37374
   339
  (Scan.optional Parse.short_ident runN -- parse_params
blanchet@37374
   340
   -- parse_relevance_override -- Scan.option Parse.nat) #>> sledgehammer_trans
blanchet@35961
   341
val parse_sledgehammer_params_command =
blanchet@35961
   342
  parse_params #>> sledgehammer_params_trans
blanchet@35961
   343
blanchet@35961
   344
val _ =
wenzelm@36970
   345
  Outer_Syntax.improper_command sledgehammerN
wenzelm@36970
   346
      "search for first-order proof using automatic theorem provers" Keyword.diag
blanchet@36394
   347
      parse_sledgehammer_command
blanchet@35961
   348
val _ =
wenzelm@36970
   349
  Outer_Syntax.command sledgehammer_paramsN
wenzelm@36970
   350
      "set and display the default parameters for Sledgehammer" Keyword.thy_decl
blanchet@36394
   351
      parse_sledgehammer_params_command
blanchet@36394
   352
blanchet@39564
   353
fun auto_sledgehammer state =
blanchet@39564
   354
  if not (!auto) then
blanchet@39564
   355
    (false, state)
blanchet@39564
   356
  else
blanchet@40250
   357
    let val ctxt = Proof.context_of state in
blanchet@40250
   358
      run_sledgehammer (get_params true ctxt []) true 1 no_relevance_override
blanchet@39564
   359
                       (minimize_command [] 1) state
blanchet@39564
   360
    end
blanchet@39564
   361
blanchet@40386
   362
val setup = Auto_Tools.register_tool (sledgehammerN, auto_sledgehammer)
blanchet@39564
   363
blanchet@35866
   364
end;