src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML
author blanchet
Tue, 23 Mar 2010 14:43:22 +0100
changeset 35963 0fce6db7babd
parent 35961 943e2582dc87
child 35964 f8c738abaed8
permissions -rw-r--r--
added a syntax for specifying facts to Sledgehammer;
e.g., sledgehammer (add: foo del: bar)
which tells the relevance filter to include "foo" but omit "bar".
blanchet@35866
     1
(*  Title:      HOL/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@35961
    11
  val default_params : theory -> (string * string) list -> params
blanchet@35961
    12
end;
blanchet@35961
    13
blanchet@35866
    14
structure Sledgehammer_Isar : sig end =
blanchet@35866
    15
struct
blanchet@35866
    16
blanchet@35866
    17
open ATP_Manager
blanchet@35866
    18
open ATP_Minimal
blanchet@35963
    19
open ATP_Wrapper
blanchet@35961
    20
open Sledgehammer_Util
blanchet@35866
    21
blanchet@35866
    22
structure K = OuterKeyword and P = OuterParse
blanchet@35866
    23
blanchet@35963
    24
datatype facta = Facta of {add: Facts.ref list, del: Facts.ref list, only: bool}
blanchet@35963
    25
blanchet@35963
    26
fun add_to_facta ns = Facta {add = ns, del = [], only = false};
blanchet@35963
    27
fun del_from_facta ns = Facta {add = [], del = ns, only = false};
blanchet@35963
    28
fun only_facta ns = Facta {add = ns, del = [], only = true};
blanchet@35963
    29
val default_facta = add_to_facta []
blanchet@35963
    30
fun merge_facta_pairwise (Facta f1) (Facta f2) =
blanchet@35963
    31
  Facta {add = #add f1 @ #add f2, del = #del f1 @ #del f2,
blanchet@35963
    32
         only = #only f1 orelse #only f2} 
blanchet@35963
    33
fun merge_facta fs = fold merge_facta_pairwise fs default_facta
blanchet@35963
    34
blanchet@35961
    35
type raw_param = string * string list
blanchet@35866
    36
blanchet@35961
    37
val default_default_params =
blanchet@35961
    38
  [("debug", "false"),
blanchet@35961
    39
   ("verbose", "false"),
blanchet@35961
    40
   ("relevance_threshold", "0.5"),
blanchet@35961
    41
   ("higher_order", "smart"),
blanchet@35961
    42
   ("respect_no_atp", "true"),
blanchet@35961
    43
   ("follow_defs", "false"),
blanchet@35961
    44
   ("isar_proof", "false"),
blanchet@35961
    45
   ("minimize_timeout", "5 s")]
blanchet@35866
    46
blanchet@35961
    47
val negated_params =
blanchet@35961
    48
  [("no_debug", "debug"),
blanchet@35961
    49
   ("quiet", "verbose"),
blanchet@35961
    50
   ("partial_types", "full_types"),
blanchet@35961
    51
   ("first_order", "higher_order"),
blanchet@35961
    52
   ("ignore_no_atp", "respect_no_atp"),
blanchet@35961
    53
   ("dont_follow_defs", "follow_defs"),
blanchet@35961
    54
   ("metis_proof", "isar_proof")]
blanchet@35866
    55
blanchet@35961
    56
val property_affected_params = ["atps", "full_types", "timeout"]
blanchet@35866
    57
blanchet@35961
    58
fun is_known_raw_param s =
blanchet@35961
    59
  AList.defined (op =) default_default_params s orelse
blanchet@35961
    60
  AList.defined (op =) negated_params s orelse
blanchet@35961
    61
  member (op =) property_affected_params s
blanchet@35866
    62
blanchet@35961
    63
fun check_raw_param (s, _) =
blanchet@35961
    64
  if is_known_raw_param s then ()
blanchet@35961
    65
  else error ("Unknown parameter: " ^ quote s ^ ".")
blanchet@35961
    66
blanchet@35961
    67
fun unnegate_raw_param (name, value) =
blanchet@35961
    68
  case AList.lookup (op =) negated_params name of
blanchet@35961
    69
    SOME name' => (name', case value of
blanchet@35961
    70
                            ["false"] => ["true"]
blanchet@35961
    71
                          | ["true"] => ["false"]
blanchet@35961
    72
                          | [] => ["false"]
blanchet@35961
    73
                          | _ => value)
blanchet@35961
    74
  | NONE => (name, value)
blanchet@35961
    75
blanchet@35961
    76
structure Data = Theory_Data(
blanchet@35961
    77
  type T = raw_param list
blanchet@35961
    78
  val empty = default_default_params |> map (apsnd single)
blanchet@35961
    79
  val extend = I
blanchet@35961
    80
  fun merge p : T = AList.merge (op =) (K true) p)
blanchet@35961
    81
blanchet@35961
    82
val set_default_raw_param = Data.map o AList.update (op =) o unnegate_raw_param
blanchet@35961
    83
fun default_raw_params thy =
blanchet@35961
    84
  [("atps", [!atps]),
blanchet@35961
    85
   ("full_types", [if !full_types then "true" else "false"]),
blanchet@35961
    86
   ("timeout", [string_of_int (!timeout) ^ " s"])] @
blanchet@35961
    87
  Data.get thy
blanchet@35961
    88
blanchet@35961
    89
val infinity_time_in_secs = 60 * 60 * 24 * 365
blanchet@35961
    90
val the_timeout = the_default (Time.fromSeconds infinity_time_in_secs)
blanchet@35961
    91
blanchet@35961
    92
fun extract_params thy default_params override_params =
blanchet@35961
    93
  let
blanchet@35961
    94
    val override_params = map unnegate_raw_param override_params
blanchet@35961
    95
    val raw_params = rev override_params @ rev default_params
blanchet@35961
    96
    val lookup = Option.map (space_implode " ") o AList.lookup (op =) raw_params
blanchet@35961
    97
    val lookup_string = the_default "" o lookup
blanchet@35961
    98
    fun general_lookup_bool option default_value name =
blanchet@35961
    99
      case lookup name of
blanchet@35961
   100
        SOME s => s |> parse_bool_option option name
blanchet@35961
   101
      | NONE => default_value
blanchet@35961
   102
    val lookup_bool = the o general_lookup_bool false (SOME false)
blanchet@35961
   103
    val lookup_bool_option = general_lookup_bool true NONE
blanchet@35961
   104
    fun lookup_time name =
blanchet@35961
   105
      the_timeout (case lookup name of
blanchet@35961
   106
                     NONE => NONE
blanchet@35961
   107
                   | SOME s => parse_time_option name s)
blanchet@35961
   108
    fun lookup_real name =
blanchet@35961
   109
      case lookup name of
blanchet@35961
   110
        NONE => 0.0
blanchet@35961
   111
      | SOME s => 0.0 (* FIXME ### *)
blanchet@35961
   112
    val debug = lookup_bool "debug"
blanchet@35961
   113
    val verbose = debug orelse lookup_bool "verbose"
blanchet@35961
   114
    val atps = lookup_string "atps" |> space_explode " "
blanchet@35961
   115
    val full_types = lookup_bool "full_types"
blanchet@35961
   116
    val relevance_threshold = lookup_real "relevance_threshold"
blanchet@35961
   117
    val higher_order = lookup_bool_option "higher_order"
blanchet@35961
   118
    val respect_no_atp = lookup_bool "respect_no_atp"
blanchet@35961
   119
    val follow_defs = lookup_bool "follow_defs"
blanchet@35961
   120
    val isar_proof = lookup_bool "isar_proof"
blanchet@35961
   121
    val timeout = lookup_time "timeout"
blanchet@35961
   122
    val minimize_timeout = lookup_time "minimize_timeout"
blanchet@35961
   123
  in
blanchet@35961
   124
    {debug = debug, verbose = verbose, atps = atps, full_types = full_types,
blanchet@35961
   125
     relevance_threshold = relevance_threshold, higher_order = higher_order,
blanchet@35961
   126
     respect_no_atp = respect_no_atp, follow_defs = follow_defs,
blanchet@35961
   127
     isar_proof = isar_proof, timeout = timeout,
blanchet@35961
   128
     minimize_timeout = minimize_timeout}
blanchet@35961
   129
  end
blanchet@35961
   130
blanchet@35963
   131
fun get_params thy = extract_params thy (default_raw_params thy)
blanchet@35963
   132
fun default_params thy = get_params thy o map (apsnd single)
blanchet@35866
   133
blanchet@35963
   134
fun atp_minimize_command override_params old_style_args fact_refs state =
blanchet@35866
   135
  let
blanchet@35961
   136
    val thy = Proof.theory_of state
blanchet@35961
   137
    val ctxt = Proof.context_of state
blanchet@35963
   138
    fun theorems_from_refs ctxt =
blanchet@35963
   139
      map (fn fact_ref =>
blanchet@35866
   140
              let
blanchet@35963
   141
                val ths = ProofContext.get_fact ctxt fact_ref
blanchet@35963
   142
                val name' = Facts.string_of_ref fact_ref
blanchet@35866
   143
              in (name', ths) end)
blanchet@35961
   144
    fun get_time_limit_arg s =
blanchet@35961
   145
      (case Int.fromString s of
blanchet@35961
   146
        SOME t => Time.fromSeconds t
blanchet@35961
   147
      | NONE => error ("Invalid time limit: " ^ quote s))
blanchet@35866
   148
    fun get_opt (name, a) (p, t) =
blanchet@35866
   149
      (case name of
blanchet@35866
   150
        "time" => (p, get_time_limit_arg a)
blanchet@35866
   151
      | "atp" => (a, t)
blanchet@35866
   152
      | n => error ("Invalid argument: " ^ n))
blanchet@35963
   153
    val {atps, minimize_timeout, ...} = get_params thy override_params
blanchet@35963
   154
    val (atp, timeout) = fold get_opt old_style_args (hd atps, minimize_timeout)
blanchet@35961
   155
    val params =
blanchet@35963
   156
      get_params thy
blanchet@35963
   157
          [("atps", [atp]),
blanchet@35963
   158
           ("minimize_timeout",
blanchet@35963
   159
            [string_of_int (Time.toSeconds timeout) ^ " s"])]
blanchet@35866
   160
    val prover =
blanchet@35961
   161
      (case get_prover thy atp of
blanchet@35866
   162
        SOME prover => prover
blanchet@35961
   163
      | NONE => error ("Unknown ATP: " ^ quote atp))
blanchet@35963
   164
    val name_thms_pairs = theorems_from_refs ctxt fact_refs
blanchet@35866
   165
  in
blanchet@35961
   166
    writeln (#2 (minimize_theorems params linear_minimize prover atp state
blanchet@35961
   167
                                   name_thms_pairs))
blanchet@35866
   168
  end
blanchet@35866
   169
blanchet@35963
   170
val runN = "run"
blanchet@35963
   171
val minimizeN = "minimize"
blanchet@35963
   172
val messagesN = "messages"
blanchet@35963
   173
val available_atpsN = "available_atps"
blanchet@35963
   174
val running_atpsN = "running_atps"
blanchet@35963
   175
val kill_atpsN = "kill_atps"
blanchet@35963
   176
val refresh_tptpN = "refresh_tptp"
blanchet@35963
   177
val helpN = "help"
blanchet@35866
   178
blanchet@35963
   179
val addN = "add"
blanchet@35963
   180
val delN = "del"
blanchet@35963
   181
val onlyN = "only"
blanchet@35963
   182
blanchet@35963
   183
fun hammer_away override_params subcommand opt_i (Facta f) state =
blanchet@35961
   184
  let
blanchet@35961
   185
    val thy = Proof.theory_of state
blanchet@35961
   186
    val _ = List.app check_raw_param override_params
blanchet@35963
   187
  in
blanchet@35963
   188
    if subcommand = runN then
blanchet@35963
   189
      sledgehammer (get_params thy override_params) (the_default 1 opt_i) state
blanchet@35963
   190
    else if subcommand = minimizeN then
blanchet@35963
   191
      atp_minimize_command override_params [] (#add f) state
blanchet@35963
   192
    else if subcommand = messagesN then
blanchet@35963
   193
      messages opt_i
blanchet@35963
   194
    else if subcommand = available_atpsN then
blanchet@35963
   195
      available_atps thy
blanchet@35963
   196
    else if subcommand = running_atpsN then
blanchet@35963
   197
      running_atps ()
blanchet@35963
   198
    else if subcommand = kill_atpsN then
blanchet@35963
   199
      kill_atps ()
blanchet@35963
   200
    else if subcommand = refresh_tptpN then
blanchet@35963
   201
      refresh_systems_on_tptp ()
blanchet@35963
   202
    else
blanchet@35963
   203
      error ("Unknown subcommand: " ^ quote subcommand ^ ".")
blanchet@35963
   204
  end
blanchet@35961
   205
blanchet@35963
   206
fun sledgehammer_trans (((params, subcommand), opt_i), facta) =
blanchet@35963
   207
  Toplevel.keep (hammer_away params subcommand opt_i facta o Toplevel.proof_of)
blanchet@35961
   208
blanchet@35961
   209
fun string_for_raw_param (name, value) = name ^ " = " ^ space_implode " " value
blanchet@35961
   210
blanchet@35963
   211
fun sledgehammer_params_trans params =
blanchet@35961
   212
  Toplevel.theory
blanchet@35963
   213
      (fold set_default_raw_param params
blanchet@35961
   214
       #> tap (fn thy => 
blanchet@35961
   215
                  writeln ("Default parameters for Sledgehammer:\n" ^
blanchet@35961
   216
                           (case rev (default_raw_params thy) of
blanchet@35961
   217
                              [] => "none"
blanchet@35961
   218
                            | params =>
blanchet@35961
   219
                              (map check_raw_param params;
blanchet@35961
   220
                               params |> map string_for_raw_param
blanchet@35961
   221
                                      |> sort_strings |> cat_lines)))))
blanchet@35961
   222
blanchet@35961
   223
val parse_key = Scan.repeat1 P.typ_group >> space_implode " "
blanchet@35961
   224
val parse_value =
blanchet@35961
   225
  Scan.repeat1 (P.minus >> single
blanchet@35961
   226
                || Scan.repeat1 (Scan.unless P.minus P.name)) >> flat
blanchet@35963
   227
val parse_param = parse_key -- Scan.optional (P.$$$ "=" |-- parse_value) []
blanchet@35963
   228
val parse_params = Scan.optional (Args.bracks (P.list parse_param)) []
blanchet@35963
   229
val parse_fact_refs =
blanchet@35963
   230
  Scan.repeat (P.xname -- Scan.option Attrib.thm_sel
blanchet@35963
   231
               >> (fn (name, interval) =>
blanchet@35963
   232
                      Facts.Named ((name, Position.none), interval)))
blanchet@35963
   233
val parse_slew =
blanchet@35963
   234
  (Args.add |-- Args.colon |-- parse_fact_refs >> add_to_facta)
blanchet@35963
   235
  || (Args.del |-- Args.colon |-- parse_fact_refs >> del_from_facta)
blanchet@35963
   236
  || (P.$$$ onlyN |-- Args.colon |-- parse_fact_refs >> only_facta)
blanchet@35963
   237
val parse_facta =
blanchet@35963
   238
  Args.parens (Scan.optional (parse_fact_refs >> only_facta) default_facta
blanchet@35963
   239
               -- Scan.repeat parse_slew) >> op :: >> merge_facta
blanchet@35961
   240
val parse_sledgehammer_command =
blanchet@35963
   241
  (parse_params -- Scan.optional P.name runN -- Scan.option P.nat
blanchet@35963
   242
   -- parse_facta)
blanchet@35963
   243
  #>> sledgehammer_trans
blanchet@35961
   244
val parse_sledgehammer_params_command =
blanchet@35961
   245
  parse_params #>> sledgehammer_params_trans
blanchet@35961
   246
blanchet@35963
   247
val parse_minimize_args =
blanchet@35963
   248
  Scan.optional (Args.bracks (P.list (P.xname --| P.$$$ "=" -- P.xname))) []
blanchet@35963
   249
blanchet@35961
   250
val _ =
blanchet@35961
   251
  OuterSyntax.improper_command "atp_kill" "kill all managed provers" K.diag
blanchet@35963
   252
    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill_atps))
blanchet@35961
   253
val _ =
blanchet@35961
   254
  OuterSyntax.improper_command "atp_info"
blanchet@35961
   255
    "print information about managed provers" K.diag
blanchet@35963
   256
    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative running_atps))
blanchet@35961
   257
val _ =
blanchet@35961
   258
  OuterSyntax.improper_command "atp_messages"
blanchet@35961
   259
    "print recent messages issued by managed provers" K.diag
blanchet@35961
   260
    (Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")") >>
blanchet@35961
   261
      (fn limit => Toplevel.no_timing
blanchet@35961
   262
                   o Toplevel.imperative (fn () => messages limit)))
blanchet@35961
   263
val _ =
blanchet@35961
   264
  OuterSyntax.improper_command "print_atps" "print external provers" K.diag
blanchet@35961
   265
    (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory o
blanchet@35963
   266
      Toplevel.keep (available_atps o Toplevel.theory_of)))
blanchet@35963
   267
val _ =
blanchet@35963
   268
  OuterSyntax.improper_command "atp_minimize"
blanchet@35963
   269
    "minimize theorem list with external prover" K.diag
blanchet@35963
   270
    (parse_minimize_args -- parse_fact_refs >> (fn (args, fact_refs) =>
blanchet@35963
   271
      Toplevel.no_timing o Toplevel.unknown_proof o
blanchet@35963
   272
        Toplevel.keep (atp_minimize_command [] args fact_refs
blanchet@35963
   273
                       o Toplevel.proof_of)))
blanchet@35963
   274
blanchet@35961
   275
val _ =
blanchet@35961
   276
  OuterSyntax.improper_command "sledgehammer"
blanchet@35961
   277
    "search for first-order proof using automatic theorem provers" K.diag
blanchet@35961
   278
    parse_sledgehammer_command
blanchet@35961
   279
val _ =
blanchet@35961
   280
  OuterSyntax.command "sledgehammer_params"
blanchet@35961
   281
    "set and display the default parameters for Sledgehammer" K.thy_decl
blanchet@35961
   282
    parse_sledgehammer_params_command
blanchet@35866
   283
blanchet@35866
   284
end;