src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML
author blanchet
Fri, 01 Jul 2011 15:53:38 +0200
changeset 44493 a867ebb12209
parent 44431 b342cd125533
child 44581 7270ae921cf2
permissions -rw-r--r--
renamed "type_sys" to "type_enc", which is more accurate
wenzelm@32564
     1
(*  Title:      HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML
wenzelm@32564
     2
    Author:     Jasmin Blanchette and Sascha Boehme and Tobias Nipkow, TU Munich
boehmes@32385
     3
*)
boehmes@32385
     4
boehmes@32385
     5
structure Mirabelle_Sledgehammer : MIRABELLE_ACTION =
boehmes@32385
     6
struct
boehmes@32385
     7
boehmes@32521
     8
val proverK = "prover"
boehmes@32541
     9
val prover_timeoutK = "prover_timeout"
boehmes@32521
    10
val keepK = "keep"
blanchet@44493
    11
val type_encK = "type_enc"
blanchet@43590
    12
val slicingK = "slicing"
blanchet@43590
    13
val e_weight_methodK = "e_weight_method"
blanchet@43590
    14
val spass_force_sosK = "spass_force_sos"
blanchet@43590
    15
val vampire_force_sosK = "vampire_force_sos"
blanchet@42623
    16
val max_relevantK = "max_relevant"
boehmes@32525
    17
val minimizeK = "minimize"
boehmes@32525
    18
val minimize_timeoutK = "minimize_timeout"
boehmes@34033
    19
val metis_ftK = "metis_ft"
blanchet@41605
    20
val reconstructorK = "reconstructor"
boehmes@32521
    21
boehmes@32521
    22
fun sh_tag id = "#" ^ string_of_int id ^ " sledgehammer: "
boehmes@32525
    23
fun minimize_tag id = "#" ^ string_of_int id ^ " minimize (sledgehammer): "
blanchet@40915
    24
fun reconstructor_tag reconstructor id =
blanchet@40915
    25
  "#" ^ string_of_int id ^ " " ^ (!reconstructor) ^ " (sledgehammer): "
boehmes@32521
    26
boehmes@32525
    27
val separator = "-----"
boehmes@32525
    28
boehmes@32521
    29
nipkow@32549
    30
datatype sh_data = ShData of {
nipkow@32549
    31
  calls: int,
nipkow@32549
    32
  success: int,
blanchet@39583
    33
  nontriv_calls: int,
blanchet@39583
    34
  nontriv_success: int,
nipkow@32585
    35
  lemmas: int,
nipkow@32810
    36
  max_lems: int,
nipkow@32549
    37
  time_isa: int,
blanchet@40243
    38
  time_prover: int,
blanchet@40243
    39
  time_prover_fail: int}
boehmes@32521
    40
blanchet@40915
    41
datatype re_data = ReData of {
nipkow@32549
    42
  calls: int,
nipkow@32549
    43
  success: int,
blanchet@39583
    44
  nontriv_calls: int,
blanchet@39583
    45
  nontriv_success: int,
nipkow@32676
    46
  proofs: int,
nipkow@32549
    47
  time: int,
nipkow@32550
    48
  timeout: int,
nipkow@32990
    49
  lemmas: int * int * int,
blanchet@39587
    50
  posns: (Position.T * bool) list
nipkow@32550
    51
  }
boehmes@32521
    52
nipkow@32571
    53
datatype min_data = MinData of {
nipkow@32609
    54
  succs: int,
blanchet@35866
    55
  ab_ratios: int
nipkow@32571
    56
  }
boehmes@32521
    57
nipkow@32810
    58
fun make_sh_data
blanchet@39583
    59
      (calls,success,nontriv_calls,nontriv_success,lemmas,max_lems,time_isa,
blanchet@40243
    60
       time_prover,time_prover_fail) =
blanchet@39583
    61
  ShData{calls=calls, success=success, nontriv_calls=nontriv_calls,
blanchet@39583
    62
         nontriv_success=nontriv_success, lemmas=lemmas, max_lems=max_lems,
blanchet@40243
    63
         time_isa=time_isa, time_prover=time_prover,
blanchet@40243
    64
         time_prover_fail=time_prover_fail}
boehmes@32521
    65
blanchet@35866
    66
fun make_min_data (succs, ab_ratios) =
blanchet@35866
    67
  MinData{succs=succs, ab_ratios=ab_ratios}
nipkow@32571
    68
blanchet@40915
    69
fun make_re_data (calls,success,nontriv_calls,nontriv_success,proofs,time,
blanchet@39583
    70
                  timeout,lemmas,posns) =
blanchet@40915
    71
  ReData{calls=calls, success=success, nontriv_calls=nontriv_calls,
blanchet@39583
    72
         nontriv_success=nontriv_success, proofs=proofs, time=time,
nipkow@32990
    73
         timeout=timeout, lemmas=lemmas, posns=posns}
boehmes@32521
    74
blanchet@39583
    75
val empty_sh_data = make_sh_data (0, 0, 0, 0, 0, 0, 0, 0, 0)
blanchet@35871
    76
val empty_min_data = make_min_data (0, 0)
blanchet@40915
    77
val empty_re_data = make_re_data (0, 0, 0, 0, 0, 0, 0, (0,0,0), [])
boehmes@32521
    78
blanchet@39583
    79
fun tuple_of_sh_data (ShData {calls, success, nontriv_calls, nontriv_success,
blanchet@39583
    80
                              lemmas, max_lems, time_isa,
blanchet@40243
    81
  time_prover, time_prover_fail}) = (calls, success, nontriv_calls,
blanchet@40243
    82
  nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail)
boehmes@32521
    83
blanchet@35866
    84
fun tuple_of_min_data (MinData {succs, ab_ratios}) = (succs, ab_ratios)
nipkow@32533
    85
blanchet@40915
    86
fun tuple_of_re_data (ReData {calls, success, nontriv_calls, nontriv_success,
blanchet@39583
    87
  proofs, time, timeout, lemmas, posns}) = (calls, success, nontriv_calls,
blanchet@39583
    88
  nontriv_success, proofs, time, timeout, lemmas, posns)
nipkow@32571
    89
boehmes@34033
    90
blanchet@40915
    91
datatype reconstructor_mode =
blanchet@40915
    92
  Unminimized | Minimized | UnminimizedFT | MinimizedFT
boehmes@34033
    93
boehmes@34033
    94
datatype data = Data of {
boehmes@34033
    95
  sh: sh_data,
boehmes@34033
    96
  min: min_data,
blanchet@40915
    97
  re_u: re_data, (* reconstructor with unminimized set of lemmas *)
blanchet@40915
    98
  re_m: re_data, (* reconstructor with minimized set of lemmas *)
blanchet@40915
    99
  re_uft: re_data, (* reconstructor with unminimized set of lemmas and fully-typed *)
blanchet@40915
   100
  re_mft: re_data, (* reconstructor with minimized set of lemmas and fully-typed *)
boehmes@34033
   101
  mini: bool   (* with minimization *)
boehmes@34033
   102
  }
boehmes@34033
   103
blanchet@40915
   104
fun make_data (sh, min, re_u, re_m, re_uft, re_mft, mini) =
blanchet@40915
   105
  Data {sh=sh, min=min, re_u=re_u, re_m=re_m, re_uft=re_uft, re_mft=re_mft,
boehmes@34033
   106
    mini=mini}
boehmes@34033
   107
boehmes@34033
   108
val empty_data = make_data (empty_sh_data, empty_min_data,
blanchet@40915
   109
  empty_re_data, empty_re_data, empty_re_data, empty_re_data, false)
boehmes@34033
   110
blanchet@40915
   111
fun map_sh_data f (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
boehmes@34033
   112
  let val sh' = make_sh_data (f (tuple_of_sh_data sh))
blanchet@40915
   113
  in make_data (sh', min, re_u, re_m, re_uft, re_mft, mini) end
boehmes@34033
   114
blanchet@40915
   115
fun map_min_data f (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
boehmes@34033
   116
  let val min' = make_min_data (f (tuple_of_min_data min))
blanchet@40915
   117
  in make_data (sh, min', re_u, re_m, re_uft, re_mft, mini) end
boehmes@34033
   118
blanchet@40915
   119
fun map_re_data f m (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
boehmes@34033
   120
  let
boehmes@34033
   121
    fun map_me g Unminimized   (u, m, uft, mft) = (g u, m, uft, mft)
boehmes@34033
   122
      | map_me g Minimized     (u, m, uft, mft) = (u, g m, uft, mft)
boehmes@34033
   123
      | map_me g UnminimizedFT (u, m, uft, mft) = (u, m, g uft, mft)
boehmes@34033
   124
      | map_me g MinimizedFT   (u, m, uft, mft) = (u, m, uft, g mft)
boehmes@34033
   125
blanchet@40915
   126
    val f' = make_re_data o f o tuple_of_re_data
boehmes@34033
   127
blanchet@40915
   128
    val (re_u', re_m', re_uft', re_mft') =
blanchet@40915
   129
      map_me f' m (re_u, re_m, re_uft, re_mft)
blanchet@40915
   130
  in make_data (sh, min, re_u', re_m', re_uft', re_mft', mini) end
boehmes@34033
   131
blanchet@40915
   132
fun set_mini mini (Data {sh, min, re_u, re_m, re_uft, re_mft, ...}) =
blanchet@40915
   133
  make_data (sh, min, re_u, re_m, re_uft, re_mft, mini)
nipkow@32990
   134
nipkow@32990
   135
fun inc_max (n:int) (s,sos,m) = (s+n, sos + n*n, Int.max(m,n));
nipkow@32536
   136
nipkow@32810
   137
val inc_sh_calls =  map_sh_data
blanchet@40243
   138
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
blanchet@40243
   139
    => (calls + 1, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail))
boehmes@32521
   140
nipkow@32810
   141
val inc_sh_success = map_sh_data
blanchet@40243
   142
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
blanchet@40243
   143
    => (calls, success + 1, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail))
blanchet@39583
   144
blanchet@39583
   145
val inc_sh_nontriv_calls =  map_sh_data
blanchet@40243
   146
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
blanchet@40243
   147
    => (calls, success, nontriv_calls + 1, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail))
blanchet@39583
   148
blanchet@39583
   149
val inc_sh_nontriv_success = map_sh_data
blanchet@40243
   150
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
blanchet@40243
   151
    => (calls, success, nontriv_calls, nontriv_success + 1, lemmas,max_lems, time_isa, time_prover, time_prover_fail))
nipkow@32585
   152
nipkow@32810
   153
fun inc_sh_lemmas n = map_sh_data
blanchet@40243
   154
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
blanchet@40243
   155
    => (calls,success,nontriv_calls, nontriv_success, lemmas+n,max_lems,time_isa,time_prover,time_prover_fail))
boehmes@32521
   156
nipkow@32810
   157
fun inc_sh_max_lems n = map_sh_data
blanchet@40243
   158
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
blanchet@40243
   159
    => (calls,success,nontriv_calls, nontriv_success, lemmas,Int.max(max_lems,n),time_isa,time_prover,time_prover_fail))
nipkow@32549
   160
nipkow@32810
   161
fun inc_sh_time_isa t = map_sh_data
blanchet@40243
   162
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
blanchet@40243
   163
    => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa + t,time_prover,time_prover_fail))
nipkow@32549
   164
blanchet@40243
   165
fun inc_sh_time_prover t = map_sh_data
blanchet@40243
   166
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
blanchet@40243
   167
    => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover + t,time_prover_fail))
nipkow@32571
   168
blanchet@40243
   169
fun inc_sh_time_prover_fail t = map_sh_data
blanchet@40243
   170
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
blanchet@40243
   171
    => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail + t))
nipkow@32571
   172
nipkow@32810
   173
val inc_min_succs = map_min_data
blanchet@35866
   174
  (fn (succs,ab_ratios) => (succs+1, ab_ratios))
nipkow@32609
   175
nipkow@32810
   176
fun inc_min_ab_ratios r = map_min_data
blanchet@35866
   177
  (fn (succs, ab_ratios) => (succs, ab_ratios+r))
nipkow@32549
   178
blanchet@40915
   179
val inc_reconstructor_calls = map_re_data
blanchet@39583
   180
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
blanchet@39583
   181
    => (calls + 1, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns))
nipkow@32549
   182
blanchet@40915
   183
val inc_reconstructor_success = map_re_data
blanchet@39583
   184
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
blanchet@39583
   185
    => (calls, success + 1, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns))
blanchet@39583
   186
blanchet@40915
   187
val inc_reconstructor_nontriv_calls = map_re_data
blanchet@39583
   188
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
blanchet@39583
   189
    => (calls, success, nontriv_calls + 1, nontriv_success, proofs, time, timeout, lemmas,posns))
blanchet@39583
   190
blanchet@40915
   191
val inc_reconstructor_nontriv_success = map_re_data
blanchet@39583
   192
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
blanchet@39583
   193
    => (calls, success, nontriv_calls, nontriv_success + 1, proofs, time, timeout, lemmas,posns))
nipkow@32676
   194
blanchet@40915
   195
val inc_reconstructor_proofs = map_re_data
blanchet@39583
   196
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
blanchet@39583
   197
    => (calls, success, nontriv_calls, nontriv_success, proofs + 1, time, timeout, lemmas,posns))
nipkow@32549
   198
blanchet@40915
   199
fun inc_reconstructor_time m t = map_re_data
blanchet@39583
   200
 (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
blanchet@39583
   201
  => (calls, success, nontriv_calls, nontriv_success, proofs, time + t, timeout, lemmas,posns)) m
nipkow@32549
   202
blanchet@40915
   203
val inc_reconstructor_timeout = map_re_data
blanchet@39583
   204
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
blanchet@39583
   205
    => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout + 1, lemmas,posns))
nipkow@32549
   206
blanchet@40915
   207
fun inc_reconstructor_lemmas m n = map_re_data
blanchet@39583
   208
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
blanchet@39583
   209
    => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, inc_max n lemmas, posns)) m
nipkow@32549
   210
blanchet@40915
   211
fun inc_reconstructor_posns m pos = map_re_data
blanchet@39583
   212
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
blanchet@39583
   213
    => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas, pos::posns)) m
boehmes@32521
   214
boehmes@32521
   215
local
boehmes@32521
   216
boehmes@32521
   217
val str = string_of_int
boehmes@32521
   218
val str3 = Real.fmt (StringCvt.FIX (SOME 3))
boehmes@32521
   219
fun percentage a b = string_of_int (a * 100 div b)
boehmes@32521
   220
fun time t = Real.fromInt t / 1000.0
boehmes@32521
   221
fun avg_time t n =
boehmes@32521
   222
  if n > 0 then (Real.fromInt t / 1000.0) / Real.fromInt n else 0.0
boehmes@32521
   223
boehmes@34033
   224
fun log_sh_data log
blanchet@40243
   225
    (calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail) =
nipkow@32810
   226
 (log ("Total number of sledgehammer calls: " ^ str calls);
nipkow@32810
   227
  log ("Number of successful sledgehammer calls: " ^ str success);
nipkow@32810
   228
  log ("Number of sledgehammer lemmas: " ^ str lemmas);
nipkow@32810
   229
  log ("Max number of sledgehammer lemmas: " ^ str max_lems);
nipkow@32810
   230
  log ("Success rate: " ^ percentage success calls ^ "%");
blanchet@39583
   231
  log ("Total number of nontrivial sledgehammer calls: " ^ str nontriv_calls);
blanchet@39583
   232
  log ("Number of successful nontrivial sledgehammer calls: " ^ str nontriv_success);
nipkow@32810
   233
  log ("Total time for sledgehammer calls (Isabelle): " ^ str3 (time time_isa));
blanchet@40243
   234
  log ("Total time for successful sledgehammer calls (ATP): " ^ str3 (time time_prover));
blanchet@40243
   235
  log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time time_prover_fail));
nipkow@32536
   236
  log ("Average time for sledgehammer calls (Isabelle): " ^
nipkow@32810
   237
    str3 (avg_time time_isa calls));
nipkow@32533
   238
  log ("Average time for successful sledgehammer calls (ATP): " ^
blanchet@40243
   239
    str3 (avg_time time_prover success));
nipkow@32536
   240
  log ("Average time for failed sledgehammer calls (ATP): " ^
blanchet@40243
   241
    str3 (avg_time time_prover_fail (calls - success)))
nipkow@32533
   242
  )
boehmes@32521
   243
nipkow@32551
   244
blanchet@39587
   245
fun str_of_pos (pos, triv) =
nipkow@32551
   246
  let val str0 = string_of_int o the_default 0
blanchet@39587
   247
  in
blanchet@39587
   248
    str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos) ^
blanchet@39587
   249
    (if triv then "[T]" else "")
blanchet@39587
   250
  end
nipkow@32551
   251
blanchet@40915
   252
fun log_re_data log tag sh_calls (re_calls, re_success, re_nontriv_calls,
blanchet@40915
   253
     re_nontriv_success, re_proofs, re_time, re_timeout,
blanchet@40915
   254
    (lemmas, lems_sos, lems_max), re_posns) =
blanchet@40915
   255
 (log ("Total number of " ^ tag ^ "reconstructor calls: " ^ str re_calls);
blanchet@40915
   256
  log ("Number of successful " ^ tag ^ "reconstructor calls: " ^ str re_success ^
blanchet@40915
   257
    " (proof: " ^ str re_proofs ^ ")");
blanchet@40915
   258
  log ("Number of " ^ tag ^ "reconstructor timeouts: " ^ str re_timeout);
blanchet@40915
   259
  log ("Success rate: " ^ percentage re_success sh_calls ^ "%");
blanchet@40915
   260
  log ("Total number of nontrivial " ^ tag ^ "reconstructor calls: " ^ str re_nontriv_calls);
blanchet@40915
   261
  log ("Number of successful nontrivial " ^ tag ^ "reconstructor calls: " ^ str re_nontriv_success ^
blanchet@40915
   262
    " (proof: " ^ str re_proofs ^ ")");
blanchet@40915
   263
  log ("Number of successful " ^ tag ^ "reconstructor lemmas: " ^ str lemmas);
blanchet@40915
   264
  log ("SOS of successful " ^ tag ^ "reconstructor lemmas: " ^ str lems_sos);
blanchet@40915
   265
  log ("Max number of successful " ^ tag ^ "reconstructor lemmas: " ^ str lems_max);
blanchet@40915
   266
  log ("Total time for successful " ^ tag ^ "reconstructor calls: " ^ str3 (time re_time));
blanchet@40915
   267
  log ("Average time for successful " ^ tag ^ "reconstructor calls: " ^
blanchet@40915
   268
    str3 (avg_time re_time re_success));
nipkow@32551
   269
  if tag=""
blanchet@40915
   270
  then log ("Proved: " ^ space_implode " " (map str_of_pos re_posns))
nipkow@32551
   271
  else ()
nipkow@32551
   272
 )
nipkow@32571
   273
blanchet@35866
   274
fun log_min_data log (succs, ab_ratios) =
nipkow@32609
   275
  (log ("Number of successful minimizations: " ^ string_of_int succs);
blanchet@35866
   276
   log ("After/before ratios: " ^ string_of_int ab_ratios)
nipkow@32571
   277
  )
nipkow@32571
   278
boehmes@32521
   279
in
boehmes@32521
   280
blanchet@40915
   281
fun log_data id log (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
boehmes@34033
   282
  let
boehmes@34033
   283
    val ShData {calls=sh_calls, ...} = sh
boehmes@34033
   284
blanchet@40915
   285
    fun app_if (ReData {calls, ...}) f = if calls > 0 then f () else ()
blanchet@40915
   286
    fun log_re tag m =
blanchet@40915
   287
      log_re_data log tag sh_calls (tuple_of_re_data m)
blanchet@40915
   288
    fun log_reconstructor (tag1, m1) (tag2, m2) = app_if m1 (fn () =>
blanchet@40915
   289
      (log_re tag1 m1; log ""; app_if m2 (fn () => log_re tag2 m2)))
boehmes@34033
   290
  in
boehmes@34033
   291
    if sh_calls > 0
boehmes@34033
   292
    then
boehmes@34033
   293
     (log ("\n\n\nReport #" ^ string_of_int id ^ ":\n");
boehmes@34033
   294
      log_sh_data log (tuple_of_sh_data sh);
boehmes@34033
   295
      log "";
boehmes@34033
   296
      if not mini
blanchet@40915
   297
      then log_reconstructor ("", re_u) ("fully-typed ", re_uft)
boehmes@34033
   298
      else
blanchet@40915
   299
        app_if re_u (fn () =>
blanchet@40915
   300
         (log_reconstructor ("unminimized ", re_u) ("unminimized fully-typed ", re_uft);
boehmes@34033
   301
          log "";
blanchet@40915
   302
          app_if re_m (fn () =>
boehmes@34033
   303
            (log_min_data log (tuple_of_min_data min); log "";
blanchet@40915
   304
             log_reconstructor ("", re_m) ("fully-typed ", re_mft))))))
boehmes@34033
   305
    else ()
boehmes@34033
   306
  end
boehmes@32521
   307
boehmes@32521
   308
end
boehmes@32521
   309
boehmes@32521
   310
boehmes@32521
   311
(* Warning: we implicitly assume single-threaded execution here! *)
wenzelm@32740
   312
val data = Unsynchronized.ref ([] : (int * data) list)
boehmes@32521
   313
wenzelm@32740
   314
fun init id thy = (Unsynchronized.change data (cons (id, empty_data)); thy)
wenzelm@32567
   315
fun done id ({log, ...}: Mirabelle.done_args) =
boehmes@32521
   316
  AList.lookup (op =) (!data) id
boehmes@32521
   317
  |> Option.map (log_data id log)
boehmes@32521
   318
  |> K ()
boehmes@32521
   319
wenzelm@32740
   320
fun change_data id f = (Unsynchronized.change data (AList.map_entry (op =) id f); ())
boehmes@32521
   321
boehmes@32521
   322
blanchet@43315
   323
fun get_prover ctxt args =
boehmes@33016
   324
  let
blanchet@40243
   325
    fun default_prover_name () =
blanchet@40250
   326
      hd (#provers (Sledgehammer_Isar.default_params ctxt []))
boehmes@33016
   327
      handle Empty => error "No ATP available."
blanchet@41335
   328
    fun get_prover name =
blanchet@43862
   329
      (name, Sledgehammer_Run.get_minimizing_prover ctxt
blanchet@43862
   330
                Sledgehammer_Provers.Normal name)
boehmes@33016
   331
  in
boehmes@33016
   332
    (case AList.lookup (op =) args proverK of
blanchet@40243
   333
      SOME name => get_prover name
blanchet@40243
   334
    | NONE => get_prover (default_prover_name ()))
boehmes@33016
   335
  end
boehmes@32525
   336
blanchet@43929
   337
type locality = ATP_Translate.locality
blanchet@38991
   338
blanchet@40915
   339
(* hack *)
blanchet@41605
   340
fun reconstructor_from_msg args msg =
blanchet@41605
   341
  (case AList.lookup (op =) args reconstructorK of
blanchet@41605
   342
    SOME name => name
blanchet@41605
   343
  | NONE =>
blanchet@44069
   344
    if String.isSubstring "metis (full_types)" msg then "metis (full_types)"
blanchet@44069
   345
    else if String.isSubstring "metis (no_types)" msg then "metis (no_types)"
blanchet@41605
   346
    else if String.isSubstring "metis" msg then "metis"
blanchet@41605
   347
    else "smt")
blanchet@40915
   348
boehmes@32521
   349
local
boehmes@32521
   350
nipkow@32536
   351
datatype sh_result =
blanchet@38991
   352
  SH_OK of int * int * (string * locality) list |
nipkow@32536
   353
  SH_FAIL of int * int |
nipkow@32536
   354
  SH_ERROR
nipkow@32536
   355
blanchet@44493
   356
fun run_sh prover_name prover type_enc max_relevant slicing e_weight_method spass_force_sos
blanchet@43590
   357
      vampire_force_sos hard_timeout timeout dir st =
boehmes@32521
   358
  let
blanchet@39242
   359
    val {context = ctxt, facts = chained_ths, goal} = Proof.goal st
blanchet@39242
   360
    val i = 1
blanchet@41585
   361
    fun change_dir (SOME dir) =
blanchet@41585
   362
        Config.put Sledgehammer_Provers.dest_dir dir
blanchet@41585
   363
        #> Config.put SMT_Config.debug_files
blanchet@43929
   364
          (dir ^ "/" ^ Name.desymbolize false (ATP_Util.timestamp ()) ^ "_"
blanchet@41586
   365
          ^ serial_string ())
blanchet@39567
   366
      | change_dir NONE = I
blanchet@39567
   367
    val st' =
blanchet@39567
   368
      st |> Proof.map_context
blanchet@39567
   369
                (change_dir dir
blanchet@43590
   370
                 #> (Option.map (Config.put ATP_Systems.e_weight_method)
blanchet@43590
   371
                       e_weight_method |> the_default I)
blanchet@43590
   372
                 #> (Option.map (Config.put ATP_Systems.spass_force_sos)
blanchet@43590
   373
                       spass_force_sos |> the_default I)
blanchet@43590
   374
                 #> (Option.map (Config.put ATP_Systems.vampire_force_sos)
blanchet@43590
   375
                       vampire_force_sos |> the_default I)
blanchet@41335
   376
                 #> Config.put Sledgehammer_Provers.measure_run_time true)
blanchet@43513
   377
    val params as {relevance_thresholds, max_relevant, slicing, ...} =
blanchet@40250
   378
      Sledgehammer_Isar.default_params ctxt
blanchet@40796
   379
          [("verbose", "true"),
blanchet@44493
   380
           ("type_enc", type_enc),
blanchet@42623
   381
           ("max_relevant", max_relevant),
blanchet@43590
   382
           ("slicing", slicing),
wenzelm@41739
   383
           ("timeout", string_of_int timeout)]
blanchet@40243
   384
    val default_max_relevant =
blanchet@43314
   385
      Sledgehammer_Provers.default_max_relevant_for_prover ctxt slicing
blanchet@43314
   386
        prover_name
blanchet@43793
   387
    val is_appropriate_prop =
blanchet@43793
   388
      Sledgehammer_Provers.is_appropriate_prop_for_prover ctxt prover_name
blanchet@40615
   389
    val is_built_in_const =
blanchet@41335
   390
      Sledgehammer_Provers.is_built_in_const_for_prover ctxt prover_name
blanchet@41189
   391
    val relevance_fudge =
blanchet@41335
   392
      Sledgehammer_Provers.relevance_fudge_for_prover ctxt prover_name
blanchet@40251
   393
    val relevance_override = {add = [], del = [], only = false}
blanchet@43929
   394
    val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal i
boehmes@32573
   395
    val time_limit =
boehmes@32573
   396
      (case hard_timeout of
boehmes@32573
   397
        NONE => I
boehmes@32573
   398
      | SOME secs => TimeLimit.timeLimit (Time.fromSeconds secs))
blanchet@43794
   399
    fun failed failure =
blanchet@43893
   400
      ({outcome = SOME failure, used_facts = [], run_time_in_msecs = NONE,
blanchet@44007
   401
        preplay = K (ATP_Reconstruct.Failed_to_Play ATP_Reconstruct.Metis),
blanchet@44102
   402
        message = K "", message_tail = ""}, ~1)
blanchet@44102
   403
    val ({outcome, used_facts, run_time_in_msecs, preplay, message,
blanchet@44102
   404
          message_tail} : Sledgehammer_Provers.prover_result,
blanchet@41523
   405
        time_isa) = time_limit (Mirabelle.cpu_time (fn () =>
blanchet@41523
   406
      let
blanchet@43794
   407
        val _ = if is_appropriate_prop concl_t then ()
blanchet@43794
   408
                else raise Fail "inappropriate"
blanchet@41523
   409
        val facts =
blanchet@44217
   410
          Sledgehammer_Filter.nearly_all_facts ctxt relevance_override
blanchet@44217
   411
                                               chained_ths hyp_ts concl_t
blanchet@44217
   412
          |> filter (is_appropriate_prop o prop_of o snd)
blanchet@44217
   413
          |> Sledgehammer_Filter.relevant_facts ctxt relevance_thresholds
blanchet@44217
   414
                 (the_default default_max_relevant max_relevant)
blanchet@44217
   415
                 is_built_in_const relevance_fudge relevance_override
blanchet@44217
   416
                 chained_ths hyp_ts concl_t
blanchet@41523
   417
        val problem =
blanchet@41523
   418
          {state = st', goal = goal, subgoal = i,
blanchet@41523
   419
           subgoal_count = Sledgehammer_Util.subgoal_count st,
blanchet@41523
   420
           facts = facts |> map Sledgehammer_Provers.Untranslated_Fact,
blanchet@42612
   421
           smt_filter = NONE}
blanchet@43892
   422
      in prover params (K (K "")) problem end)) ()
blanchet@43794
   423
      handle TimeLimit.TimeOut => failed ATP_Proof.TimedOut
blanchet@43794
   424
           | Fail "inappropriate" => failed ATP_Proof.Inappropriate
blanchet@40620
   425
    val time_prover = run_time_in_msecs |> the_default ~1
blanchet@44102
   426
    val msg = message (preplay ()) ^ message_tail
boehmes@32521
   427
  in
blanchet@36405
   428
    case outcome of
blanchet@43893
   429
      NONE => (msg, SH_OK (time_isa, time_prover, used_facts))
blanchet@43893
   430
    | SOME _ => (msg, SH_FAIL (time_isa, time_prover))
boehmes@32521
   431
  end
blanchet@38228
   432
  handle ERROR msg => ("error: " ^ msg, SH_ERROR)
boehmes@32521
   433
boehmes@32454
   434
fun thms_of_name ctxt name =
boehmes@32454
   435
  let
wenzelm@36970
   436
    val lex = Keyword.get_lexicons
wenzelm@43232
   437
    val get = maps (Proof_Context.get_fact ctxt o fst)
boehmes@32454
   438
  in
boehmes@32454
   439
    Source.of_string name
wenzelm@40772
   440
    |> Symbol.source
wenzelm@36969
   441
    |> Token.source {do_recover=SOME false} lex Position.start
wenzelm@36969
   442
    |> Token.source_proper
wenzelm@36969
   443
    |> Source.source Token.stopper (Parse_Spec.xthms1 >> get) NONE
boehmes@32454
   444
    |> Source.exhaust
boehmes@32454
   445
  end
boehmes@32452
   446
boehmes@32498
   447
in
boehmes@32498
   448
blanchet@40915
   449
fun run_sledgehammer trivial args reconstructor named_thms id ({pre=st, log, ...}: Mirabelle.run_args) =
boehmes@32385
   450
  let
blanchet@39586
   451
    val triv_str = if trivial then "[T] " else ""
nipkow@32536
   452
    val _ = change_data id inc_sh_calls
blanchet@39583
   453
    val _ = if trivial then () else change_data id inc_sh_nontriv_calls
blanchet@43315
   454
    val (prover_name, prover) = get_prover (Proof.context_of st) args
blanchet@44493
   455
    val type_enc = AList.lookup (op =) args type_encK |> the_default "smart"
blanchet@42623
   456
    val max_relevant = AList.lookup (op =) args max_relevantK |> the_default "smart"
blanchet@43590
   457
    val slicing = AList.lookup (op =) args slicingK |> the_default "true"
blanchet@43590
   458
    val e_weight_method = AList.lookup (op =) args e_weight_methodK
blanchet@43590
   459
    val spass_force_sos = AList.lookup (op =) args spass_force_sosK
blanchet@43590
   460
      |> Option.map (curry (op <>) "false")
blanchet@43590
   461
    val vampire_force_sos = AList.lookup (op =) args vampire_force_sosK
blanchet@43590
   462
      |> Option.map (curry (op <>) "false")
boehmes@32525
   463
    val dir = AList.lookup (op =) args keepK
boehmes@32541
   464
    val timeout = Mirabelle.get_int_setting args (prover_timeoutK, 30)
blanchet@41516
   465
    (* always use a hard timeout, but give some slack so that the automatic
blanchet@41516
   466
       minimizer has a chance to do its magic *)
blanchet@41516
   467
    val hard_timeout = SOME (2 * timeout)
blanchet@41403
   468
    val (msg, result) =
blanchet@44493
   469
      run_sh prover_name prover type_enc max_relevant slicing e_weight_method spass_force_sos
blanchet@43590
   470
        vampire_force_sos hard_timeout timeout dir st
boehmes@32525
   471
  in
nipkow@32536
   472
    case result of
blanchet@40243
   473
      SH_OK (time_isa, time_prover, names) =>
blanchet@38939
   474
        let
blanchet@43929
   475
          fun get_thms (_, ATP_Translate.Chained) = NONE
blanchet@39061
   476
            | get_thms (name, loc) =
blanchet@39061
   477
              SOME ((name, loc), thms_of_name (Proof.context_of st) name)
boehmes@32525
   478
        in
nipkow@32810
   479
          change_data id inc_sh_success;
blanchet@39583
   480
          if trivial then () else change_data id inc_sh_nontriv_success;
nipkow@32810
   481
          change_data id (inc_sh_lemmas (length names));
nipkow@32810
   482
          change_data id (inc_sh_max_lems (length names));
nipkow@32810
   483
          change_data id (inc_sh_time_isa time_isa);
blanchet@40243
   484
          change_data id (inc_sh_time_prover time_prover);
blanchet@41605
   485
          reconstructor := reconstructor_from_msg args msg;
blanchet@39061
   486
          named_thms := SOME (map_filter get_thms names);
blanchet@39586
   487
          log (sh_tag id ^ triv_str ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
blanchet@40243
   488
            string_of_int time_prover ^ ") [" ^ prover_name ^ "]:\n" ^ msg)
boehmes@32525
   489
        end
blanchet@40243
   490
    | SH_FAIL (time_isa, time_prover) =>
nipkow@32536
   491
        let
nipkow@32536
   492
          val _ = change_data id (inc_sh_time_isa time_isa)
blanchet@40243
   493
          val _ = change_data id (inc_sh_time_prover_fail time_prover)
blanchet@39586
   494
        in log (sh_tag id ^ triv_str ^ "failed: " ^ msg) end
nipkow@32536
   495
    | SH_ERROR => log (sh_tag id ^ "failed: " ^ msg)
boehmes@32525
   496
  end
boehmes@32525
   497
boehmes@32525
   498
end
boehmes@32525
   499
blanchet@40915
   500
fun run_minimize args reconstructor named_thms id
blanchet@40915
   501
        ({pre=st, log, ...}: Mirabelle.run_args) =
boehmes@32525
   502
  let
blanchet@40250
   503
    val ctxt = Proof.context_of st
nipkow@32571
   504
    val n0 = length (these (!named_thms))
blanchet@43315
   505
    val (prover_name, _) = get_prover ctxt args
blanchet@44493
   506
    val type_enc = AList.lookup (op =) args type_encK |> the_default "smart"
boehmes@32525
   507
    val timeout =
boehmes@32525
   508
      AList.lookup (op =) args minimize_timeoutK
wenzelm@40875
   509
      |> Option.map (fst o read_int o raw_explode)  (* FIXME Symbol.explode (?) *)
boehmes@32525
   510
      |> the_default 5
blanchet@43905
   511
    val params = Sledgehammer_Isar.default_params ctxt
blanchet@41403
   512
      [("provers", prover_name),
blanchet@41403
   513
       ("verbose", "true"),
blanchet@44493
   514
       ("type_enc", type_enc),
wenzelm@41739
   515
       ("timeout", string_of_int timeout)]
blanchet@37587
   516
    val minimize =
blanchet@42613
   517
      Sledgehammer_Minimize.minimize_facts prover_name params
blanchet@43905
   518
          true 1 (Sledgehammer_Util.subgoal_count st)
boehmes@32525
   519
    val _ = log separator
blanchet@44102
   520
    val (used_facts, (preplay, message, message_tail)) =
blanchet@44102
   521
      minimize st (these (!named_thms))
blanchet@44102
   522
    val msg = message (preplay ()) ^ message_tail
boehmes@32525
   523
  in
blanchet@43893
   524
    case used_facts of
blanchet@43893
   525
      SOME named_thms' =>
nipkow@32609
   526
        (change_data id inc_min_succs;
nipkow@32609
   527
         change_data id (inc_min_ab_ratios ((100 * length named_thms') div n0));
nipkow@32571
   528
         if length named_thms' = n0
nipkow@32571
   529
         then log (minimize_tag id ^ "already minimal")
blanchet@41605
   530
         else (reconstructor := reconstructor_from_msg args msg;
blanchet@40915
   531
               named_thms := SOME named_thms';
nipkow@32571
   532
               log (minimize_tag id ^ "succeeded:\n" ^ msg))
nipkow@32571
   533
        )
blanchet@43893
   534
    | NONE => log (minimize_tag id ^ "failed: " ^ msg)
boehmes@32525
   535
  end
boehmes@32525
   536
boehmes@32525
   537
blanchet@40915
   538
fun run_reconstructor trivial full m name reconstructor named_thms id
wenzelm@32567
   539
    ({pre=st, timeout, log, pos, ...}: Mirabelle.run_args) =
boehmes@32525
   540
  let
blanchet@40915
   541
    fun do_reconstructor thms ctxt =
blanchet@41605
   542
      (if !reconstructor = "sledgehammer_tac" then
blanchet@41605
   543
         (fn ctxt => fn thms =>
blanchet@41605
   544
            Method.insert_tac thms THEN'
blanchet@41605
   545
            Sledgehammer_Tactics.sledgehammer_as_unsound_oracle_tac ctxt)
blanchet@41605
   546
       else if !reconstructor = "smt" then
blanchet@41402
   547
         SMT_Solver.smt_tac
blanchet@44069
   548
       else if full orelse !reconstructor = "metis (full_types)" then
blanchet@44493
   549
         Metis_Tactics.metis_tac [Metis_Tactics.full_type_enc]
blanchet@44069
   550
       else if !reconstructor = "metis (no_types)" then
blanchet@44493
   551
         Metis_Tactics.metis_tac [Metis_Tactics.no_type_enc]
blanchet@41402
   552
       else
blanchet@44053
   553
         Metis_Tactics.metis_tac []) ctxt thms
blanchet@40915
   554
    fun apply_reconstructor thms =
blanchet@40915
   555
      Mirabelle.can_apply timeout (do_reconstructor thms) st
boehmes@32521
   556
boehmes@32521
   557
    fun with_time (false, t) = "failed (" ^ string_of_int t ^ ")"
blanchet@40915
   558
      | with_time (true, t) = (change_data id (inc_reconstructor_success m);
blanchet@40915
   559
          if trivial then ()
blanchet@40915
   560
          else change_data id (inc_reconstructor_nontriv_success m);
blanchet@40915
   561
          change_data id (inc_reconstructor_lemmas m (length named_thms));
blanchet@40915
   562
          change_data id (inc_reconstructor_time m t);
blanchet@40915
   563
          change_data id (inc_reconstructor_posns m (pos, trivial));
blanchet@40915
   564
          if name = "proof" then change_data id (inc_reconstructor_proofs m)
blanchet@40915
   565
          else ();
boehmes@32521
   566
          "succeeded (" ^ string_of_int t ^ ")")
blanchet@40915
   567
    fun timed_reconstructor thms =
blanchet@40915
   568
      (with_time (Mirabelle.cpu_time apply_reconstructor thms), true)
blanchet@40915
   569
      handle TimeLimit.TimeOut => (change_data id (inc_reconstructor_timeout m);
boehmes@34048
   570
               ("timeout", false))
boehmes@34048
   571
           | ERROR msg => ("error: " ^ msg, false)
boehmes@32521
   572
boehmes@32525
   573
    val _ = log separator
blanchet@40915
   574
    val _ = change_data id (inc_reconstructor_calls m)
blanchet@40915
   575
    val _ = if trivial then ()
blanchet@40915
   576
            else change_data id (inc_reconstructor_nontriv_calls m)
boehmes@32521
   577
  in
boehmes@32525
   578
    maps snd named_thms
blanchet@40915
   579
    |> timed_reconstructor
blanchet@40915
   580
    |>> log o prefix (reconstructor_tag reconstructor id)
boehmes@34048
   581
    |> snd
boehmes@32521
   582
  end
boehmes@32432
   583
blanchet@41524
   584
val try_timeout = seconds 5.0
blanchet@39583
   585
boehmes@34033
   586
fun sledgehammer_action args id (st as {pre, name, ...}: Mirabelle.run_args) =
wenzelm@35596
   587
  let val goal = Thm.major_prem_of (#goal (Proof.goal pre)) in
wenzelm@35596
   588
    if can Logic.dest_conjunction goal orelse can Logic.dest_equals goal
wenzelm@35596
   589
    then () else
wenzelm@35596
   590
    let
blanchet@40915
   591
      val reconstructor = Unsynchronized.ref ""
blanchet@38939
   592
      val named_thms =
blanchet@38991
   593
        Unsynchronized.ref (NONE : ((string * locality) * thm list) list option)
wenzelm@35596
   594
      val minimize = AList.defined (op =) args minimizeK
wenzelm@35596
   595
      val metis_ft = AList.defined (op =) args metis_ftK
blanchet@43861
   596
      val trivial =
blanchet@43861
   597
        Try_Methods.try_methods (SOME try_timeout) ([], [], [], []) pre
blanchet@41524
   598
        handle TimeLimit.TimeOut => false
blanchet@40915
   599
      fun apply_reconstructor m1 m2 =
wenzelm@35596
   600
        if metis_ft
wenzelm@35596
   601
        then
blanchet@40915
   602
          if not (Mirabelle.catch_result (reconstructor_tag reconstructor) false
blanchet@40915
   603
              (run_reconstructor trivial false m1 name reconstructor
blanchet@40915
   604
                   (these (!named_thms))) id st)
wenzelm@35596
   605
          then
blanchet@40915
   606
            (Mirabelle.catch_result (reconstructor_tag reconstructor) false
blanchet@40915
   607
              (run_reconstructor trivial true m2 name reconstructor
blanchet@40915
   608
                   (these (!named_thms))) id st; ())
wenzelm@35596
   609
          else ()
wenzelm@35596
   610
        else
blanchet@40915
   611
          (Mirabelle.catch_result (reconstructor_tag reconstructor) false
blanchet@40915
   612
            (run_reconstructor trivial false m1 name reconstructor
blanchet@40915
   613
                 (these (!named_thms))) id st; ())
wenzelm@35596
   614
    in 
wenzelm@35596
   615
      change_data id (set_mini minimize);
blanchet@40915
   616
      Mirabelle.catch sh_tag (run_sledgehammer trivial args reconstructor
blanchet@40915
   617
                                               named_thms) id st;
wenzelm@35596
   618
      if is_some (!named_thms)
boehmes@34033
   619
      then
blanchet@40915
   620
       (apply_reconstructor Unminimized UnminimizedFT;
wenzelm@35596
   621
        if minimize andalso not (null (these (!named_thms)))
boehmes@34048
   622
        then
blanchet@40915
   623
         (Mirabelle.catch minimize_tag
blanchet@40915
   624
              (run_minimize args reconstructor named_thms) id st;
blanchet@40915
   625
          apply_reconstructor Minimized MinimizedFT)
wenzelm@35596
   626
        else ())
wenzelm@35596
   627
      else ()
wenzelm@35596
   628
    end
nipkow@32810
   629
  end
boehmes@32385
   630
boehmes@32511
   631
fun invoke args =
blanchet@44431
   632
  Mirabelle.register (init, sledgehammer_action args, done)
boehmes@32385
   633
boehmes@32385
   634
end