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