src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML
author blanchet
Sun, 25 Apr 2010 14:40:36 +0200
changeset 36394 1a48d18449d8
parent 36378 f32c567dbcaa
child 36395 e73923451f6f
permissions -rw-r--r--
move "neg_clausify" method and "clausify" attribute to "sledgehammer_isar.ML"
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@35961
     9
  type params = ATP_Manager.params
blanchet@35961
    10
blanchet@36371
    11
  val atps: string Unsynchronized.ref
blanchet@36371
    12
  val timeout: int Unsynchronized.ref
blanchet@36371
    13
  val full_types: bool Unsynchronized.ref
blanchet@35961
    14
  val default_params : theory -> (string * string) list -> params
blanchet@36394
    15
  val setup: theory -> theory
blanchet@35961
    16
end;
blanchet@35961
    17
blanchet@35969
    18
structure Sledgehammer_Isar : SLEDGEHAMMER_ISAR =
blanchet@35866
    19
struct
blanchet@35866
    20
blanchet@35969
    21
open Sledgehammer_Util
blanchet@36394
    22
open Sledgehammer_Fact_Preprocessor
blanchet@35866
    23
open ATP_Manager
blanchet@36376
    24
open ATP_Systems
blanchet@36375
    25
open Sledgehammer_Fact_Minimizer
blanchet@35866
    26
blanchet@35866
    27
structure K = OuterKeyword and P = OuterParse
blanchet@35866
    28
blanchet@36394
    29
(** Proof method used in Isar proofs **)
blanchet@36394
    30
blanchet@36394
    31
val neg_clausify_setup =
blanchet@36394
    32
  Method.setup @{binding neg_clausify}
blanchet@36394
    33
               (Scan.succeed (SIMPLE_METHOD' o neg_clausify_tac))
blanchet@36394
    34
               "conversion of goal to negated conjecture clauses"
blanchet@36394
    35
blanchet@36394
    36
val parse_clausify_attribute =
blanchet@36394
    37
  Scan.lift OuterParse.nat
blanchet@36394
    38
  >> (fn i => Thm.rule_attribute (fn context => fn th =>
blanchet@36394
    39
                  let val thy = Context.theory_of context in
blanchet@36394
    40
                    Meson.make_meta_clause (nth (cnf_axiom thy th) i)
blanchet@36394
    41
                  end))
blanchet@36394
    42
blanchet@36394
    43
(** Attribute for converting a theorem into clauses **)
blanchet@36394
    44
blanchet@36394
    45
val clausify_setup =
blanchet@36394
    46
  Attrib.setup @{binding clausify} parse_clausify_attribute
blanchet@36394
    47
               "conversion of theorem to clauses"
blanchet@36394
    48
blanchet@36394
    49
(** Sledgehammer commands **)
blanchet@36394
    50
blanchet@35964
    51
fun add_to_relevance_override ns : relevance_override =
blanchet@35964
    52
  {add = ns, del = [], only = false}
blanchet@35964
    53
fun del_from_relevance_override ns : relevance_override =
blanchet@35964
    54
  {add = [], del = ns, only = false}
blanchet@35964
    55
fun only_relevance_override ns : relevance_override =
blanchet@35964
    56
  {add = ns, del = [], only = true}
blanchet@36188
    57
val no_relevance_override = add_to_relevance_override []
blanchet@36003
    58
fun merge_relevance_override_pairwise (r1 : relevance_override)
blanchet@36003
    59
                                      (r2 : relevance_override) =
blanchet@35964
    60
  {add = #add r1 @ #add r2, del = #del r1 @ #del r2,
blanchet@36183
    61
   only = #only r1 andalso #only r2}
blanchet@36188
    62
fun merge_relevance_overrides rs =
blanchet@36188
    63
  fold merge_relevance_override_pairwise rs (only_relevance_override [])
blanchet@35963
    64
blanchet@36371
    65
(*** parameters ***)
blanchet@36371
    66
blanchet@36371
    67
val atps = Unsynchronized.ref (default_atps_param_value ())
blanchet@36371
    68
val timeout = Unsynchronized.ref 60
blanchet@36371
    69
val full_types = Unsynchronized.ref false
blanchet@36371
    70
blanchet@36371
    71
val _ =
blanchet@36371
    72
  ProofGeneralPgip.add_preference Preferences.category_proof
blanchet@36371
    73
      (Preferences.string_pref atps
blanchet@36371
    74
          "Sledgehammer: ATPs"
blanchet@36371
    75
          "Default automatic provers (separated by whitespace)")
blanchet@36371
    76
blanchet@36371
    77
val _ =
blanchet@36371
    78
  ProofGeneralPgip.add_preference Preferences.category_proof
blanchet@36371
    79
      (Preferences.int_pref timeout
blanchet@36373
    80
          "Sledgehammer: Time Limit"
blanchet@36371
    81
          "ATPs will be interrupted after this time (in seconds)")
blanchet@36371
    82
blanchet@36371
    83
val _ =
blanchet@36371
    84
  ProofGeneralPgip.add_preference Preferences.category_proof
blanchet@36371
    85
      (Preferences.bool_pref full_types
blanchet@36373
    86
          "Sledgehammer: Full Types" "ATPs will use full type information")
blanchet@36371
    87
blanchet@35961
    88
type raw_param = string * string list
blanchet@35866
    89
blanchet@35961
    90
val default_default_params =
blanchet@35961
    91
  [("debug", "false"),
blanchet@35961
    92
   ("verbose", "false"),
blanchet@36141
    93
   ("overlord", "false"),
blanchet@36235
    94
   ("explicit_apply", "false"),
blanchet@36058
    95
   ("respect_no_atp", "true"),
blanchet@35964
    96
   ("relevance_threshold", "50"),
blanchet@36058
    97
   ("convergence", "320"),
blanchet@36220
    98
   ("theory_relevant", "smart"),
blanchet@35961
    99
   ("higher_order", "smart"),
blanchet@35961
   100
   ("follow_defs", "false"),
blanchet@35961
   101
   ("isar_proof", "false"),
blanchet@36064
   102
   ("modulus", "1"),
blanchet@36064
   103
   ("sorts", "false"),
blanchet@35961
   104
   ("minimize_timeout", "5 s")]
blanchet@35866
   105
blanchet@36138
   106
val alias_params =
blanchet@36138
   107
  [("atp", "atps")]
blanchet@36138
   108
val negated_alias_params =
blanchet@35961
   109
  [("no_debug", "debug"),
blanchet@35961
   110
   ("quiet", "verbose"),
blanchet@36141
   111
   ("no_overlord", "overlord"),
blanchet@36235
   112
   ("implicit_apply", "explicit_apply"),
blanchet@36058
   113
   ("ignore_no_atp", "respect_no_atp"),
blanchet@35961
   114
   ("partial_types", "full_types"),
blanchet@36220
   115
   ("theory_irrelevant", "theory_relevant"),
blanchet@35961
   116
   ("first_order", "higher_order"),
blanchet@35961
   117
   ("dont_follow_defs", "follow_defs"),
blanchet@36064
   118
   ("metis_proof", "isar_proof"),
blanchet@36064
   119
   ("no_sorts", "sorts")]
blanchet@35866
   120
blanchet@36281
   121
val params_for_minimize =
blanchet@36281
   122
  ["full_types", "explicit_apply", "higher_order", "isar_proof", "modulus",
blanchet@36281
   123
   "sorts", "minimize_timeout"]
blanchet@36281
   124
blanchet@35969
   125
val property_dependent_params = ["atps", "full_types", "timeout"]
blanchet@35866
   126
blanchet@35961
   127
fun is_known_raw_param s =
blanchet@35961
   128
  AList.defined (op =) default_default_params s orelse
blanchet@36138
   129
  AList.defined (op =) alias_params s orelse
blanchet@36138
   130
  AList.defined (op =) negated_alias_params s orelse
blanchet@35969
   131
  member (op =) property_dependent_params s
blanchet@35866
   132
blanchet@35961
   133
fun check_raw_param (s, _) =
blanchet@35961
   134
  if is_known_raw_param s then ()
blanchet@35961
   135
  else error ("Unknown parameter: " ^ quote s ^ ".")
blanchet@35961
   136
blanchet@36138
   137
fun unalias_raw_param (name, value) =
blanchet@36138
   138
  case AList.lookup (op =) alias_params name of
blanchet@36138
   139
    SOME name' => (name', value)
blanchet@36138
   140
  | NONE =>
blanchet@36138
   141
    case AList.lookup (op =) negated_alias_params name of
blanchet@36138
   142
      SOME name' => (name', case value of
blanchet@36138
   143
                              ["false"] => ["true"]
blanchet@36138
   144
                            | ["true"] => ["false"]
blanchet@36138
   145
                            | [] => ["false"]
blanchet@36138
   146
                            | _ => value)
blanchet@36138
   147
    | NONE => (name, value)
blanchet@35961
   148
blanchet@35961
   149
structure Data = Theory_Data(
blanchet@35961
   150
  type T = raw_param list
blanchet@35961
   151
  val empty = default_default_params |> map (apsnd single)
blanchet@35961
   152
  val extend = I
blanchet@35961
   153
  fun merge p : T = AList.merge (op =) (K true) p)
blanchet@35961
   154
blanchet@36138
   155
val set_default_raw_param = Data.map o AList.update (op =) o unalias_raw_param
blanchet@35961
   156
fun default_raw_params thy =
blanchet@35961
   157
  Data.get thy
blanchet@36138
   158
  |> fold (AList.default (op =))
blanchet@36290
   159
          [("atps", [!atps]),
blanchet@36138
   160
           ("full_types", [if !full_types then "true" else "false"]),
blanchet@36140
   161
           ("timeout", let val timeout = !timeout in
blanchet@36140
   162
                         [if timeout <= 0 then "none"
blanchet@36140
   163
                          else string_of_int timeout ^ " s"]
blanchet@36140
   164
                       end)]
blanchet@35961
   165
blanchet@35961
   166
val infinity_time_in_secs = 60 * 60 * 24 * 365
blanchet@35961
   167
val the_timeout = the_default (Time.fromSeconds infinity_time_in_secs)
blanchet@35961
   168
blanchet@35961
   169
fun extract_params thy default_params override_params =
blanchet@35961
   170
  let
blanchet@36138
   171
    val override_params = map unalias_raw_param override_params
blanchet@35961
   172
    val raw_params = rev override_params @ rev default_params
blanchet@35961
   173
    val lookup = Option.map (space_implode " ") o AList.lookup (op =) raw_params
blanchet@35961
   174
    val lookup_string = the_default "" o lookup
blanchet@35961
   175
    fun general_lookup_bool option default_value name =
blanchet@35961
   176
      case lookup name of
blanchet@35968
   177
        SOME s => parse_bool_option option name s
blanchet@35961
   178
      | NONE => default_value
blanchet@35961
   179
    val lookup_bool = the o general_lookup_bool false (SOME false)
blanchet@35961
   180
    val lookup_bool_option = general_lookup_bool true NONE
blanchet@35961
   181
    fun lookup_time name =
blanchet@35961
   182
      the_timeout (case lookup name of
blanchet@35961
   183
                     NONE => NONE
blanchet@35961
   184
                   | SOME s => parse_time_option name s)
blanchet@35964
   185
    fun lookup_int name =
blanchet@35961
   186
      case lookup name of
blanchet@35964
   187
        NONE => 0
blanchet@35964
   188
      | SOME s => case Int.fromString s of
blanchet@35964
   189
                    SOME n => n
blanchet@35964
   190
                  | NONE => error ("Parameter " ^ quote name ^
blanchet@35964
   191
                                   " must be assigned an integer value.")
blanchet@35961
   192
    val debug = lookup_bool "debug"
blanchet@35961
   193
    val verbose = debug orelse lookup_bool "verbose"
blanchet@36141
   194
    val overlord = lookup_bool "overlord"
blanchet@35961
   195
    val atps = lookup_string "atps" |> space_explode " "
blanchet@35961
   196
    val full_types = lookup_bool "full_types"
blanchet@36235
   197
    val explicit_apply = lookup_bool "explicit_apply"
blanchet@36058
   198
    val respect_no_atp = lookup_bool "respect_no_atp"
blanchet@35964
   199
    val relevance_threshold =
blanchet@35964
   200
      0.01 * Real.fromInt (lookup_int "relevance_threshold")
blanchet@36058
   201
    val convergence = 0.01 * Real.fromInt (lookup_int "convergence")
blanchet@36220
   202
    val theory_relevant = lookup_bool_option "theory_relevant"
blanchet@35961
   203
    val higher_order = lookup_bool_option "higher_order"
blanchet@35961
   204
    val follow_defs = lookup_bool "follow_defs"
blanchet@35961
   205
    val isar_proof = lookup_bool "isar_proof"
blanchet@36064
   206
    val modulus = Int.max (1, lookup_int "modulus")
blanchet@36064
   207
    val sorts = lookup_bool "sorts"
blanchet@35961
   208
    val timeout = lookup_time "timeout"
blanchet@35961
   209
    val minimize_timeout = lookup_time "minimize_timeout"
blanchet@35961
   210
  in
blanchet@36141
   211
    {debug = debug, verbose = verbose, overlord = overlord, atps = atps,
blanchet@36235
   212
     full_types = full_types, explicit_apply = explicit_apply,
blanchet@36235
   213
     respect_no_atp = respect_no_atp, relevance_threshold = relevance_threshold,
blanchet@36235
   214
     convergence = convergence, theory_relevant = theory_relevant,
blanchet@36235
   215
     higher_order = higher_order, follow_defs = follow_defs,
blanchet@36235
   216
     isar_proof = isar_proof, modulus = modulus, sorts = sorts,
blanchet@36235
   217
     timeout = timeout, minimize_timeout = minimize_timeout}
blanchet@35961
   218
  end
blanchet@35961
   219
blanchet@35963
   220
fun get_params thy = extract_params thy (default_raw_params thy)
blanchet@35963
   221
fun default_params thy = get_params thy o map (apsnd single)
blanchet@35866
   222
blanchet@36373
   223
(* Sledgehammer the given subgoal *)
blanchet@36373
   224
blanchet@36373
   225
fun run {atps = [], ...} _ _ _ _ = error "No ATP is set."
blanchet@36373
   226
  | run (params as {atps, timeout, ...}) i relevance_override minimize_command
blanchet@36373
   227
        proof_state =
blanchet@36373
   228
    let
blanchet@36373
   229
      val birth_time = Time.now ()
blanchet@36373
   230
      val death_time = Time.+ (birth_time, timeout)
blanchet@36373
   231
      val _ = kill_atps ()  (* race w.r.t. other invocations of Sledgehammer *)
blanchet@36373
   232
      val _ = priority "Sledgehammering..."
blanchet@36373
   233
      val _ = List.app (start_prover_thread params birth_time death_time i
blanchet@36373
   234
                                            relevance_override minimize_command
blanchet@36373
   235
                                            proof_state) atps
blanchet@36373
   236
    in () end
blanchet@36373
   237
blanchet@36378
   238
fun minimize override_params i fact_refs state =
blanchet@35866
   239
  let
blanchet@35961
   240
    val thy = Proof.theory_of state
blanchet@35961
   241
    val ctxt = Proof.context_of state
blanchet@36378
   242
    val theorems_from_refs =
blanchet@36378
   243
      map o pairf Facts.string_of_ref o ProofContext.get_fact
blanchet@35963
   244
    val name_thms_pairs = theorems_from_refs ctxt fact_refs
blanchet@35866
   245
  in
blanchet@36378
   246
    priority (#2 (minimize_theorems (get_params thy override_params) i state
blanchet@36378
   247
                                    name_thms_pairs))
blanchet@35866
   248
  end
blanchet@35866
   249
blanchet@36373
   250
val sledgehammerN = "sledgehammer"
blanchet@36373
   251
val sledgehammer_paramsN = "sledgehammer_params"
blanchet@36373
   252
blanchet@35963
   253
val runN = "run"
blanchet@35963
   254
val minimizeN = "minimize"
blanchet@35963
   255
val messagesN = "messages"
blanchet@35963
   256
val available_atpsN = "available_atps"
blanchet@35963
   257
val running_atpsN = "running_atps"
blanchet@35963
   258
val kill_atpsN = "kill_atps"
blanchet@35963
   259
val refresh_tptpN = "refresh_tptp"
blanchet@35963
   260
val helpN = "help"
blanchet@35866
   261
blanchet@36139
   262
fun minimizize_raw_param_name "timeout" = "minimize_timeout"
blanchet@36139
   263
  | minimizize_raw_param_name name = name
blanchet@36139
   264
blanchet@36281
   265
val is_raw_param_relevant_for_minimize =
blanchet@36281
   266
  member (op =) params_for_minimize o fst o unalias_raw_param
blanchet@36281
   267
fun string_for_raw_param (key, values) =
blanchet@36281
   268
  key ^ (case space_implode " " values of
blanchet@36281
   269
           "" => ""
blanchet@36281
   270
         | value => " = " ^ value)
blanchet@36281
   271
blanchet@36281
   272
fun minimize_command override_params i atp_name facts =
blanchet@36373
   273
  sledgehammerN ^ " " ^ minimizeN ^ " [atp = " ^ atp_name ^
blanchet@36281
   274
  (override_params |> filter is_raw_param_relevant_for_minimize
blanchet@36281
   275
                   |> implode o map (prefix ", " o string_for_raw_param)) ^
blanchet@36281
   276
  "] (" ^ space_implode " " facts ^ ")" ^
blanchet@36281
   277
  (if i = 1 then "" else " " ^ string_of_int i)
blanchet@36281
   278
blanchet@35964
   279
fun hammer_away override_params subcommand opt_i relevance_override state =
blanchet@35961
   280
  let
blanchet@35961
   281
    val thy = Proof.theory_of state
blanchet@35961
   282
    val _ = List.app check_raw_param override_params
blanchet@35963
   283
  in
blanchet@35963
   284
    if subcommand = runN then
blanchet@36281
   285
      let val i = the_default 1 opt_i in
blanchet@36373
   286
        run (get_params thy override_params) i relevance_override
blanchet@36373
   287
            (minimize_command override_params i) state
blanchet@36281
   288
      end
blanchet@35963
   289
    else if subcommand = minimizeN then
blanchet@36378
   290
      minimize (map (apfst minimizize_raw_param_name) override_params)
blanchet@36139
   291
               (the_default 1 opt_i) (#add relevance_override) state
blanchet@35963
   292
    else if subcommand = messagesN then
blanchet@35963
   293
      messages opt_i
blanchet@35963
   294
    else if subcommand = available_atpsN then
blanchet@35963
   295
      available_atps thy
blanchet@35963
   296
    else if subcommand = running_atpsN then
blanchet@35963
   297
      running_atps ()
blanchet@35963
   298
    else if subcommand = kill_atpsN then
blanchet@35963
   299
      kill_atps ()
blanchet@35963
   300
    else if subcommand = refresh_tptpN then
blanchet@35963
   301
      refresh_systems_on_tptp ()
blanchet@35963
   302
    else
blanchet@35963
   303
      error ("Unknown subcommand: " ^ quote subcommand ^ ".")
blanchet@35963
   304
  end
blanchet@35961
   305
blanchet@35964
   306
fun sledgehammer_trans (((subcommand, params), relevance_override), opt_i) =
blanchet@35964
   307
  Toplevel.keep (hammer_away params subcommand opt_i relevance_override
blanchet@35964
   308
                 o Toplevel.proof_of)
blanchet@35961
   309
blanchet@35961
   310
fun string_for_raw_param (name, value) = name ^ " = " ^ space_implode " " value
blanchet@35961
   311
blanchet@35963
   312
fun sledgehammer_params_trans params =
blanchet@35961
   313
  Toplevel.theory
blanchet@35963
   314
      (fold set_default_raw_param params
blanchet@35961
   315
       #> tap (fn thy => 
blanchet@35961
   316
                  writeln ("Default parameters for Sledgehammer:\n" ^
blanchet@35961
   317
                           (case rev (default_raw_params thy) of
blanchet@35961
   318
                              [] => "none"
blanchet@35961
   319
                            | params =>
blanchet@35961
   320
                              (map check_raw_param params;
blanchet@35961
   321
                               params |> map string_for_raw_param
blanchet@35961
   322
                                      |> sort_strings |> cat_lines)))))
blanchet@35961
   323
blanchet@35961
   324
val parse_key = Scan.repeat1 P.typ_group >> space_implode " "
blanchet@35964
   325
val parse_value = Scan.repeat1 P.xname
blanchet@35963
   326
val parse_param = parse_key -- Scan.optional (P.$$$ "=" |-- parse_value) []
blanchet@35963
   327
val parse_params = Scan.optional (Args.bracks (P.list parse_param)) []
blanchet@35963
   328
val parse_fact_refs =
blanchet@35964
   329
  Scan.repeat1 (Scan.unless (P.name -- Args.colon)
blanchet@35964
   330
                            (P.xname -- Scan.option Attrib.thm_sel)
blanchet@35964
   331
                >> (fn (name, interval) =>
blanchet@35964
   332
                       Facts.Named ((name, Position.none), interval)))
blanchet@35964
   333
val parse_relevance_chunk =
blanchet@35964
   334
  (Args.add |-- Args.colon |-- parse_fact_refs >> add_to_relevance_override)
blanchet@35964
   335
  || (Args.del |-- Args.colon |-- parse_fact_refs
blanchet@35964
   336
      >> del_from_relevance_override)
blanchet@36188
   337
  || (parse_fact_refs >> only_relevance_override)
blanchet@35964
   338
val parse_relevance_override =
blanchet@36188
   339
  Scan.optional (Args.parens (Scan.repeat parse_relevance_chunk
blanchet@36188
   340
                              >> merge_relevance_overrides))
blanchet@36188
   341
                (add_to_relevance_override [])
blanchet@35961
   342
val parse_sledgehammer_command =
blanchet@35964
   343
  (Scan.optional P.short_ident runN -- parse_params -- parse_relevance_override
blanchet@35964
   344
   -- Scan.option P.nat) #>> sledgehammer_trans
blanchet@35961
   345
val parse_sledgehammer_params_command =
blanchet@35961
   346
  parse_params #>> sledgehammer_params_trans
blanchet@35961
   347
blanchet@35961
   348
val _ =
blanchet@36373
   349
  OuterSyntax.improper_command sledgehammerN
blanchet@36394
   350
      "search for first-order proof using automatic theorem provers" K.diag
blanchet@36394
   351
      parse_sledgehammer_command
blanchet@35961
   352
val _ =
blanchet@36373
   353
  OuterSyntax.command sledgehammer_paramsN
blanchet@36394
   354
      "set and display the default parameters for Sledgehammer" K.thy_decl
blanchet@36394
   355
      parse_sledgehammer_params_command
blanchet@36394
   356
blanchet@36394
   357
val setup =
blanchet@36394
   358
  neg_clausify_setup
blanchet@36394
   359
  #> clausify_setup
blanchet@35866
   360
blanchet@35866
   361
end;