src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML
author blanchet
Fri, 16 Apr 2010 16:13:49 +0200
changeset 36183 f723348b2231
parent 36141 6490319b1703
child 36188 35b9f0db49a0
permissions -rw-r--r--
Sledgehammer: the empty set of fact () should mean nothing, not unchanged
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@35969
    14
structure Sledgehammer_Isar : SLEDGEHAMMER_ISAR =
blanchet@35866
    15
struct
blanchet@35866
    16
blanchet@35969
    17
open Sledgehammer_Util
blanchet@35866
    18
open ATP_Manager
blanchet@35866
    19
open ATP_Minimal
blanchet@35963
    20
open ATP_Wrapper
blanchet@35866
    21
blanchet@35866
    22
structure K = OuterKeyword and P = OuterParse
blanchet@35866
    23
blanchet@35964
    24
fun add_to_relevance_override ns : relevance_override =
blanchet@35964
    25
  {add = ns, del = [], only = false}
blanchet@35964
    26
fun del_from_relevance_override ns : relevance_override =
blanchet@35964
    27
  {add = [], del = ns, only = false}
blanchet@35964
    28
fun only_relevance_override ns : relevance_override =
blanchet@35964
    29
  {add = ns, del = [], only = true}
blanchet@36003
    30
fun merge_relevance_override_pairwise (r1 : relevance_override)
blanchet@36003
    31
                                      (r2 : relevance_override) =
blanchet@35964
    32
  {add = #add r1 @ #add r2, del = #del r1 @ #del r2,
blanchet@36183
    33
   only = #only r1 andalso #only r2}
blanchet@36183
    34
fun merge_relevance_overrides (r, rs) =
blanchet@36183
    35
  fold merge_relevance_override_pairwise rs r
blanchet@35963
    36
blanchet@35961
    37
type raw_param = string * string list
blanchet@35866
    38
blanchet@35961
    39
val default_default_params =
blanchet@35961
    40
  [("debug", "false"),
blanchet@35961
    41
   ("verbose", "false"),
blanchet@36141
    42
   ("overlord", "false"),
blanchet@36058
    43
   ("respect_no_atp", "true"),
blanchet@35964
    44
   ("relevance_threshold", "50"),
blanchet@36058
    45
   ("convergence", "320"),
blanchet@36059
    46
   ("theory_const", "smart"),
blanchet@35961
    47
   ("higher_order", "smart"),
blanchet@35961
    48
   ("follow_defs", "false"),
blanchet@35961
    49
   ("isar_proof", "false"),
blanchet@36064
    50
   ("modulus", "1"),
blanchet@36064
    51
   ("sorts", "false"),
blanchet@35961
    52
   ("minimize_timeout", "5 s")]
blanchet@35866
    53
blanchet@36138
    54
val alias_params =
blanchet@36138
    55
  [("atp", "atps")]
blanchet@36138
    56
val negated_alias_params =
blanchet@35961
    57
  [("no_debug", "debug"),
blanchet@35961
    58
   ("quiet", "verbose"),
blanchet@36141
    59
   ("no_overlord", "overlord"),
blanchet@36058
    60
   ("ignore_no_atp", "respect_no_atp"),
blanchet@35961
    61
   ("partial_types", "full_types"),
blanchet@36059
    62
   ("no_theory_const", "theory_const"),
blanchet@35961
    63
   ("first_order", "higher_order"),
blanchet@35961
    64
   ("dont_follow_defs", "follow_defs"),
blanchet@36064
    65
   ("metis_proof", "isar_proof"),
blanchet@36064
    66
   ("no_sorts", "sorts")]
blanchet@35866
    67
blanchet@35969
    68
val property_dependent_params = ["atps", "full_types", "timeout"]
blanchet@35866
    69
blanchet@35961
    70
fun is_known_raw_param s =
blanchet@35961
    71
  AList.defined (op =) default_default_params s orelse
blanchet@36138
    72
  AList.defined (op =) alias_params s orelse
blanchet@36138
    73
  AList.defined (op =) negated_alias_params s orelse
blanchet@35969
    74
  member (op =) property_dependent_params s
blanchet@35866
    75
blanchet@35961
    76
fun check_raw_param (s, _) =
blanchet@35961
    77
  if is_known_raw_param s then ()
blanchet@35961
    78
  else error ("Unknown parameter: " ^ quote s ^ ".")
blanchet@35961
    79
blanchet@36138
    80
fun unalias_raw_param (name, value) =
blanchet@36138
    81
  case AList.lookup (op =) alias_params name of
blanchet@36138
    82
    SOME name' => (name', value)
blanchet@36138
    83
  | NONE =>
blanchet@36138
    84
    case AList.lookup (op =) negated_alias_params name of
blanchet@36138
    85
      SOME name' => (name', case value of
blanchet@36138
    86
                              ["false"] => ["true"]
blanchet@36138
    87
                            | ["true"] => ["false"]
blanchet@36138
    88
                            | [] => ["false"]
blanchet@36138
    89
                            | _ => value)
blanchet@36138
    90
    | NONE => (name, value)
blanchet@35961
    91
blanchet@35961
    92
structure Data = Theory_Data(
blanchet@35961
    93
  type T = raw_param list
blanchet@35961
    94
  val empty = default_default_params |> map (apsnd single)
blanchet@35961
    95
  val extend = I
blanchet@35961
    96
  fun merge p : T = AList.merge (op =) (K true) p)
blanchet@35961
    97
blanchet@36138
    98
val set_default_raw_param = Data.map o AList.update (op =) o unalias_raw_param
blanchet@35961
    99
fun default_raw_params thy =
blanchet@35961
   100
  Data.get thy
blanchet@36138
   101
  |> fold (AList.default (op =))
blanchet@36138
   102
          [("atps", [!atps]),
blanchet@36138
   103
           ("full_types", [if !full_types then "true" else "false"]),
blanchet@36140
   104
           ("timeout", let val timeout = !timeout in
blanchet@36140
   105
                         [if timeout <= 0 then "none"
blanchet@36140
   106
                          else string_of_int timeout ^ " s"]
blanchet@36140
   107
                       end)]
blanchet@35961
   108
blanchet@35961
   109
val infinity_time_in_secs = 60 * 60 * 24 * 365
blanchet@35961
   110
val the_timeout = the_default (Time.fromSeconds infinity_time_in_secs)
blanchet@35961
   111
blanchet@35961
   112
fun extract_params thy default_params override_params =
blanchet@35961
   113
  let
blanchet@36138
   114
    val override_params = map unalias_raw_param override_params
blanchet@35961
   115
    val raw_params = rev override_params @ rev default_params
blanchet@35961
   116
    val lookup = Option.map (space_implode " ") o AList.lookup (op =) raw_params
blanchet@35961
   117
    val lookup_string = the_default "" o lookup
blanchet@35961
   118
    fun general_lookup_bool option default_value name =
blanchet@35961
   119
      case lookup name of
blanchet@35968
   120
        SOME s => parse_bool_option option name s
blanchet@35961
   121
      | NONE => default_value
blanchet@35961
   122
    val lookup_bool = the o general_lookup_bool false (SOME false)
blanchet@35961
   123
    val lookup_bool_option = general_lookup_bool true NONE
blanchet@35961
   124
    fun lookup_time name =
blanchet@35961
   125
      the_timeout (case lookup name of
blanchet@35961
   126
                     NONE => NONE
blanchet@35961
   127
                   | SOME s => parse_time_option name s)
blanchet@35964
   128
    fun lookup_int name =
blanchet@35961
   129
      case lookup name of
blanchet@35964
   130
        NONE => 0
blanchet@35964
   131
      | SOME s => case Int.fromString s of
blanchet@35964
   132
                    SOME n => n
blanchet@35964
   133
                  | NONE => error ("Parameter " ^ quote name ^
blanchet@35964
   134
                                   " must be assigned an integer value.")
blanchet@35961
   135
    val debug = lookup_bool "debug"
blanchet@35961
   136
    val verbose = debug orelse lookup_bool "verbose"
blanchet@36141
   137
    val overlord = lookup_bool "overlord"
blanchet@35961
   138
    val atps = lookup_string "atps" |> space_explode " "
blanchet@35961
   139
    val full_types = lookup_bool "full_types"
blanchet@36058
   140
    val respect_no_atp = lookup_bool "respect_no_atp"
blanchet@35964
   141
    val relevance_threshold =
blanchet@35964
   142
      0.01 * Real.fromInt (lookup_int "relevance_threshold")
blanchet@36058
   143
    val convergence = 0.01 * Real.fromInt (lookup_int "convergence")
blanchet@36059
   144
    val theory_const = lookup_bool_option "theory_const"
blanchet@35961
   145
    val higher_order = lookup_bool_option "higher_order"
blanchet@35961
   146
    val follow_defs = lookup_bool "follow_defs"
blanchet@35961
   147
    val isar_proof = lookup_bool "isar_proof"
blanchet@36064
   148
    val modulus = Int.max (1, lookup_int "modulus")
blanchet@36064
   149
    val sorts = lookup_bool "sorts"
blanchet@35961
   150
    val timeout = lookup_time "timeout"
blanchet@35961
   151
    val minimize_timeout = lookup_time "minimize_timeout"
blanchet@35961
   152
  in
blanchet@36141
   153
    {debug = debug, verbose = verbose, overlord = overlord, atps = atps,
blanchet@36141
   154
     full_types = full_types, respect_no_atp = respect_no_atp,
blanchet@36141
   155
     relevance_threshold = relevance_threshold, convergence = convergence,
blanchet@36141
   156
     theory_const = theory_const, higher_order = higher_order,
blanchet@36141
   157
     follow_defs = follow_defs, isar_proof = isar_proof, modulus = modulus,
blanchet@36141
   158
     sorts = sorts, timeout = timeout, minimize_timeout = minimize_timeout}
blanchet@35961
   159
  end
blanchet@35961
   160
blanchet@35963
   161
fun get_params thy = extract_params thy (default_raw_params thy)
blanchet@35963
   162
fun default_params thy = get_params thy o map (apsnd single)
blanchet@35866
   163
blanchet@36139
   164
fun minimize override_params old_style_args i fact_refs state =
blanchet@35866
   165
  let
blanchet@35961
   166
    val thy = Proof.theory_of state
blanchet@35961
   167
    val ctxt = Proof.context_of state
blanchet@35963
   168
    fun theorems_from_refs ctxt =
blanchet@35963
   169
      map (fn fact_ref =>
blanchet@35866
   170
              let
blanchet@35963
   171
                val ths = ProofContext.get_fact ctxt fact_ref
blanchet@35963
   172
                val name' = Facts.string_of_ref fact_ref
blanchet@35866
   173
              in (name', ths) end)
blanchet@35961
   174
    fun get_time_limit_arg s =
blanchet@35961
   175
      (case Int.fromString s of
blanchet@35961
   176
        SOME t => Time.fromSeconds t
blanchet@35961
   177
      | NONE => error ("Invalid time limit: " ^ quote s))
blanchet@35866
   178
    fun get_opt (name, a) (p, t) =
blanchet@35866
   179
      (case name of
blanchet@35866
   180
        "time" => (p, get_time_limit_arg a)
blanchet@35866
   181
      | "atp" => (a, t)
blanchet@35866
   182
      | n => error ("Invalid argument: " ^ n))
blanchet@35963
   183
    val {atps, minimize_timeout, ...} = get_params thy override_params
blanchet@35963
   184
    val (atp, timeout) = fold get_opt old_style_args (hd atps, minimize_timeout)
blanchet@35961
   185
    val params =
blanchet@35963
   186
      get_params thy
blanchet@35963
   187
          [("atps", [atp]),
blanchet@35963
   188
           ("minimize_timeout",
blanchet@36140
   189
            [string_of_int (Time.toMilliseconds timeout) ^ " ms"])]
blanchet@35866
   190
    val prover =
blanchet@35961
   191
      (case get_prover thy atp of
blanchet@35866
   192
        SOME prover => prover
blanchet@35961
   193
      | NONE => error ("Unknown ATP: " ^ quote atp))
blanchet@35963
   194
    val name_thms_pairs = theorems_from_refs ctxt fact_refs
blanchet@35866
   195
  in
blanchet@36063
   196
    writeln (#2 (minimize_theorems params linear_minimize prover atp i state
blanchet@35961
   197
                                   name_thms_pairs))
blanchet@35866
   198
  end
blanchet@35866
   199
blanchet@35963
   200
val runN = "run"
blanchet@35963
   201
val minimizeN = "minimize"
blanchet@35963
   202
val messagesN = "messages"
blanchet@35963
   203
val available_atpsN = "available_atps"
blanchet@35963
   204
val running_atpsN = "running_atps"
blanchet@35963
   205
val kill_atpsN = "kill_atps"
blanchet@35963
   206
val refresh_tptpN = "refresh_tptp"
blanchet@35963
   207
val helpN = "help"
blanchet@35866
   208
blanchet@35963
   209
val addN = "add"
blanchet@35963
   210
val delN = "del"
blanchet@35963
   211
val onlyN = "only"
blanchet@35963
   212
blanchet@36139
   213
fun minimizize_raw_param_name "timeout" = "minimize_timeout"
blanchet@36139
   214
  | minimizize_raw_param_name name = name
blanchet@36139
   215
blanchet@35964
   216
fun hammer_away override_params subcommand opt_i relevance_override state =
blanchet@35961
   217
  let
blanchet@35961
   218
    val thy = Proof.theory_of state
blanchet@35961
   219
    val _ = List.app check_raw_param override_params
blanchet@35963
   220
  in
blanchet@35963
   221
    if subcommand = runN then
blanchet@35964
   222
      sledgehammer (get_params thy override_params) (the_default 1 opt_i)
blanchet@35964
   223
                   relevance_override state
blanchet@35963
   224
    else if subcommand = minimizeN then
blanchet@36139
   225
      minimize (map (apfst minimizize_raw_param_name) override_params) []
blanchet@36139
   226
               (the_default 1 opt_i) (#add relevance_override) state
blanchet@35963
   227
    else if subcommand = messagesN then
blanchet@35963
   228
      messages opt_i
blanchet@35963
   229
    else if subcommand = available_atpsN then
blanchet@35963
   230
      available_atps thy
blanchet@35963
   231
    else if subcommand = running_atpsN then
blanchet@35963
   232
      running_atps ()
blanchet@35963
   233
    else if subcommand = kill_atpsN then
blanchet@35963
   234
      kill_atps ()
blanchet@35963
   235
    else if subcommand = refresh_tptpN then
blanchet@35963
   236
      refresh_systems_on_tptp ()
blanchet@35963
   237
    else
blanchet@35963
   238
      error ("Unknown subcommand: " ^ quote subcommand ^ ".")
blanchet@35963
   239
  end
blanchet@35961
   240
blanchet@35964
   241
fun sledgehammer_trans (((subcommand, params), relevance_override), opt_i) =
blanchet@35964
   242
  Toplevel.keep (hammer_away params subcommand opt_i relevance_override
blanchet@35964
   243
                 o Toplevel.proof_of)
blanchet@35961
   244
blanchet@35961
   245
fun string_for_raw_param (name, value) = name ^ " = " ^ space_implode " " value
blanchet@35961
   246
blanchet@35963
   247
fun sledgehammer_params_trans params =
blanchet@35961
   248
  Toplevel.theory
blanchet@35963
   249
      (fold set_default_raw_param params
blanchet@35961
   250
       #> tap (fn thy => 
blanchet@35961
   251
                  writeln ("Default parameters for Sledgehammer:\n" ^
blanchet@35961
   252
                           (case rev (default_raw_params thy) of
blanchet@35961
   253
                              [] => "none"
blanchet@35961
   254
                            | params =>
blanchet@35961
   255
                              (map check_raw_param params;
blanchet@35961
   256
                               params |> map string_for_raw_param
blanchet@35961
   257
                                      |> sort_strings |> cat_lines)))))
blanchet@35961
   258
blanchet@35961
   259
val parse_key = Scan.repeat1 P.typ_group >> space_implode " "
blanchet@35964
   260
val parse_value = Scan.repeat1 P.xname
blanchet@35963
   261
val parse_param = parse_key -- Scan.optional (P.$$$ "=" |-- parse_value) []
blanchet@35963
   262
val parse_params = Scan.optional (Args.bracks (P.list parse_param)) []
blanchet@35963
   263
val parse_fact_refs =
blanchet@35964
   264
  Scan.repeat1 (Scan.unless (P.name -- Args.colon)
blanchet@35964
   265
                            (P.xname -- Scan.option Attrib.thm_sel)
blanchet@35964
   266
                >> (fn (name, interval) =>
blanchet@35964
   267
                       Facts.Named ((name, Position.none), interval)))
blanchet@35964
   268
val parse_relevance_chunk =
blanchet@35964
   269
  (Args.add |-- Args.colon |-- parse_fact_refs >> add_to_relevance_override)
blanchet@35964
   270
  || (Args.del |-- Args.colon |-- parse_fact_refs
blanchet@35964
   271
      >> del_from_relevance_override)
blanchet@35964
   272
  || (P.$$$ onlyN |-- Args.colon |-- parse_fact_refs >> only_relevance_override)
blanchet@35964
   273
val parse_relevance_override =
blanchet@35964
   274
  Scan.optional (Args.parens
blanchet@35964
   275
      (Scan.optional (parse_fact_refs >> only_relevance_override)
blanchet@36183
   276
                     (only_relevance_override [])
blanchet@35964
   277
       -- Scan.repeat parse_relevance_chunk)
blanchet@36183
   278
       >> merge_relevance_overrides)
blanchet@36183
   279
      (add_to_relevance_override [])
blanchet@35961
   280
val parse_sledgehammer_command =
blanchet@35964
   281
  (Scan.optional P.short_ident runN -- parse_params -- parse_relevance_override
blanchet@35964
   282
   -- Scan.option P.nat) #>> sledgehammer_trans
blanchet@35961
   283
val parse_sledgehammer_params_command =
blanchet@35961
   284
  parse_params #>> sledgehammer_params_trans
blanchet@35961
   285
blanchet@35963
   286
val parse_minimize_args =
blanchet@35964
   287
  Scan.optional (Args.bracks (P.list (P.short_ident --| P.$$$ "=" -- P.xname)))
blanchet@35964
   288
                []
blanchet@35961
   289
val _ =
blanchet@35961
   290
  OuterSyntax.improper_command "atp_kill" "kill all managed provers" K.diag
blanchet@35963
   291
    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill_atps))
blanchet@35961
   292
val _ =
blanchet@35961
   293
  OuterSyntax.improper_command "atp_info"
blanchet@35961
   294
    "print information about managed provers" K.diag
blanchet@35963
   295
    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative running_atps))
blanchet@35961
   296
val _ =
blanchet@35961
   297
  OuterSyntax.improper_command "atp_messages"
blanchet@35961
   298
    "print recent messages issued by managed provers" K.diag
blanchet@35961
   299
    (Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")") >>
blanchet@35961
   300
      (fn limit => Toplevel.no_timing
blanchet@35961
   301
                   o Toplevel.imperative (fn () => messages limit)))
blanchet@35961
   302
val _ =
blanchet@35961
   303
  OuterSyntax.improper_command "print_atps" "print external provers" K.diag
blanchet@35961
   304
    (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory o
blanchet@35963
   305
      Toplevel.keep (available_atps o Toplevel.theory_of)))
blanchet@35963
   306
val _ =
blanchet@35963
   307
  OuterSyntax.improper_command "atp_minimize"
blanchet@35963
   308
    "minimize theorem list with external prover" K.diag
blanchet@35963
   309
    (parse_minimize_args -- parse_fact_refs >> (fn (args, fact_refs) =>
blanchet@35963
   310
      Toplevel.no_timing o Toplevel.unknown_proof o
blanchet@36139
   311
        Toplevel.keep (minimize [] args 1 fact_refs o Toplevel.proof_of)))
blanchet@35963
   312
blanchet@35961
   313
val _ =
blanchet@35961
   314
  OuterSyntax.improper_command "sledgehammer"
blanchet@35961
   315
    "search for first-order proof using automatic theorem provers" K.diag
blanchet@35961
   316
    parse_sledgehammer_command
blanchet@35961
   317
val _ =
blanchet@35961
   318
  OuterSyntax.command "sledgehammer_params"
blanchet@35961
   319
    "set and display the default parameters for Sledgehammer" K.thy_decl
blanchet@35961
   320
    parse_sledgehammer_params_command
blanchet@35866
   321
blanchet@35866
   322
end;